diff --git a/examples/in_progress/mini-compiler/compiler.mlw b/examples/in_progress/mini-compiler/compiler.mlw
index 47f5b7743cfe3c51153c34b35eaf43f3b1d6171b..617b2a32387a606596544d6eca79e2fb06c17d16 100644
--- a/examples/in_progress/mini-compiler/compiler.mlw
+++ b/examples/in_progress/mini-compiler/compiler.mlw
@@ -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
 
 
diff --git a/examples/in_progress/mini-compiler/compiler/compiler_Compile_com_WP_parameter_compile_com_natural_1.v b/examples/in_progress/mini-compiler/compiler/compiler_Compile_com_WP_parameter_compile_com_natural_1.v
deleted file mode 100644
index f7233ca94b84e133e9b2391a0d45a4ef3975aace..0000000000000000000000000000000000000000
--- a/examples/in_progress/mini-compiler/compiler/compiler_Compile_com_WP_parameter_compile_com_natural_1.v
+++ /dev/null
@@ -1,905 +0,0 @@
-(* 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.
-
-(* 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 *)
-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))).
-
-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 *)
-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)).
-
-Parameter iconst_post: forall {a:Type} {a_WT:WhyType a}, Z -> (func a (func Z
-  (func machine_state (func machine_state bool)))).
-
-Axiom iconst_post_def : forall {a:Type} {a_WT:WhyType a}, forall (n:Z) (x:a)
-  (p:Z) (ms:machine_state) (ms':machine_state),
-  ((infix_at (infix_at (infix_at (infix_at (iconst_post n: (func a (func Z
-  (func machine_state (func machine_state bool))))) x) p) ms)
-  ms') = true) <-> forall (s:(list Z)) (m:(map id Z)), (ms = (VMS p s m)) ->
-  (ms' = (VMS (p + 1%Z)%Z (Init.Datatypes.cons n s) m)).
-
-Parameter ivar_post: forall {a:Type} {a_WT:WhyType a}, id -> (func a (func Z
-  (func machine_state (func machine_state bool)))).
-
-Axiom ivar_post_def : forall {a:Type} {a_WT:WhyType a}, forall (x:id) (a1:a)
-  (p:Z) (ms:machine_state) (ms':machine_state),
-  ((infix_at (infix_at (infix_at (infix_at (ivar_post x: (func a (func Z
-  (func machine_state (func machine_state bool))))) a1) p) ms)
-  ms') = true) <-> forall (s:(list Z)) (m:(map id Z)), (ms = (VMS p s m)) ->
-  (ms' = (VMS (p + 1%Z)%Z (Init.Datatypes.cons (get m x) s) m)).
-
-(* Why3 assumption *)
-Definition binop := (func Z (func Z Z)).
-
-Parameter ibinop_pre: forall {a:Type} {a_WT:WhyType a}, (func a (func Z (func
-  machine_state bool))).
-
-Axiom ibinop_pre_def : forall {a:Type} {a_WT:WhyType a}, forall (x:a) (p:Z)
-  (ms:machine_state), ((infix_at (infix_at (infix_at (ibinop_pre : (func a
-  (func Z (func machine_state bool)))) x) p) ms) = true) <-> exists n1:Z,
-  exists n2:Z, exists s:(list Z), exists m:(map id Z), (ms = (VMS p
-  (Init.Datatypes.cons n2 (Init.Datatypes.cons n1 s)) m)).
-
-Axiom ibinop_pre_lemma : forall {a:Type} {a_WT:WhyType a}, forall (x:a) (p:Z)
-  (ms:machine_state), ((infix_at (infix_at (infix_at (ibinop_pre : (func a
-  (func Z (func machine_state bool)))) x) p) ms) = true) <-> exists n1:Z,
-  exists n2:Z, exists s:(list Z), exists m:(map id Z), (ms = (VMS p
-  (Init.Datatypes.cons n2 (Init.Datatypes.cons n1 s)) m)).
-
-Parameter ibinop_post: forall {a:Type} {a_WT:WhyType a}, (func Z (func Z
-  Z)) -> (func a (func Z (func machine_state (func machine_state bool)))).
-
-Axiom ibinop_post_def : forall {a:Type} {a_WT:WhyType a}, forall (op:(func Z
-  (func Z Z))) (x:a) (p:Z) (ms:machine_state) (ms':machine_state),
-  ((infix_at (infix_at (infix_at (infix_at (ibinop_post op: (func a (func Z
-  (func machine_state (func machine_state bool))))) x) p) ms)
-  ms') = true) <-> forall (n1:Z) (n2:Z) (s:(list Z)) (m:(map id Z)),
-  (ms = (VMS p (Init.Datatypes.cons n2 (Init.Datatypes.cons n1 s)) m)) ->
-  (ms' = (VMS (p + 1%Z)%Z (Init.Datatypes.cons (infix_at (infix_at op n1)
-  n2) s) m)).
-
-Axiom ibinop_post_lemma : forall {a:Type} {a_WT:WhyType a}, forall (op:(func
-  Z (func Z Z))) (x:a) (p:Z) (ms:machine_state) (ms':machine_state),
-  ((infix_at (infix_at (infix_at (infix_at (ibinop_post op: (func a (func Z
-  (func machine_state (func machine_state bool))))) x) p) ms)
-  ms') = true) <-> forall (n1:Z) (n2:Z) (s:(list Z)) (m:(map id Z)),
-  (ms = (VMS p (Init.Datatypes.cons n2 (Init.Datatypes.cons n1 s)) m)) ->
-  (ms' = (VMS (p + 1%Z)%Z (Init.Datatypes.cons (infix_at (infix_at op n1)
-  n2) s) m)).
-
-Parameter plus: (func Z (func Z Z)).
-
-Axiom plus_def : forall (x:Z) (y:Z), ((infix_at (infix_at plus x)
-  y) = (x + y)%Z).
-
-Parameter sub: (func Z (func Z Z)).
-
-Axiom sub_def : forall (x:Z) (y:Z), ((infix_at (infix_at sub x)
-  y) = (x - y)%Z).
-
-Parameter mul: (func Z (func Z Z)).
-
-Axiom mul_def : forall (x:Z) (y:Z), ((infix_at (infix_at mul x)
-  y) = (x * y)%Z).
-
-Parameter inil_post: forall {a:Type} {a_WT:WhyType a}, (func a (func Z (func
-  machine_state (func machine_state bool)))).
-
-Axiom inil_post_def : forall {a:Type} {a_WT:WhyType a}, forall (x:a) (p:Z)
-  (ms:machine_state) (ms':machine_state),
-  ((infix_at (infix_at (infix_at (infix_at (inil_post : (func a (func Z (func
-  machine_state (func machine_state bool))))) x) p) ms) ms') = true) <->
-  (ms = ms').
-
-Parameter ibranch_post: forall {a:Type} {a_WT:WhyType a}, Z -> (func a (func
-  Z (func machine_state (func machine_state bool)))).
-
-Axiom ibranch_post_def : forall {a:Type} {a_WT:WhyType a}, forall (ofs:Z)
-  (x:a) (p:Z) (ms:machine_state) (ms':machine_state),
-  ((infix_at (infix_at (infix_at (infix_at (ibranch_post ofs: (func a (func Z
-  (func machine_state (func machine_state bool))))) x) p) ms)
-  ms') = true) <-> forall (s:(list Z)) (m:(map id Z)), (ms = (VMS p s m)) ->
-  (ms' = (VMS ((p + 1%Z)%Z + ofs)%Z s m)).
-
-(* Why3 assumption *)
-Definition cond := (func Z (func Z bool)).
-
-Parameter icjump_post: forall {a:Type} {a_WT:WhyType a}, (func Z (func Z
-  bool)) -> Z -> (func a (func Z (func machine_state (func machine_state
-  bool)))).
-
-Axiom icjump_post_def : forall {a:Type} {a_WT:WhyType a}, forall (cond1:(func
-  Z (func Z bool))) (ofs:Z) (x:a) (p:Z) (ms:machine_state)
-  (ms':machine_state),
-  ((infix_at (infix_at (infix_at (infix_at (icjump_post cond1 ofs: (func a
-  (func Z (func machine_state (func machine_state bool))))) x) p) ms)
-  ms') = true) <-> forall (n1:Z) (n2:Z) (s:(list Z)) (m:(map id Z)),
-  (ms = (VMS p (Init.Datatypes.cons n2 (Init.Datatypes.cons n1 s)) m)) ->
-  ((((infix_at (infix_at cond1 n1) n2) = true) ->
-  (ms' = (VMS ((p + ofs)%Z + 1%Z)%Z s m))) /\ ((~ ((infix_at (infix_at cond1
-  n1) n2) = true)) -> (ms' = (VMS (p + 1%Z)%Z s m)))).
-
-Axiom icjump_post_lemma : forall {a:Type} {a_WT:WhyType a},
-  forall (cond1:(func Z (func Z bool))) (ofs:Z) (x:a) (p:Z)
-  (ms:machine_state) (ms':machine_state),
-  ((infix_at (infix_at (infix_at (infix_at (icjump_post cond1 ofs: (func a
-  (func Z (func machine_state (func machine_state bool))))) x) p) ms)
-  ms') = true) <-> forall (n1:Z) (n2:Z) (s:(list Z)) (m:(map id Z)),
-  (ms = (VMS p (Init.Datatypes.cons n2 (Init.Datatypes.cons n1 s)) m)) ->
-  ((((infix_at (infix_at cond1 n1) n2) = true) ->
-  (ms' = (VMS ((p + ofs)%Z + 1%Z)%Z s m))) /\ ((~ ((infix_at (infix_at cond1
-  n1) n2) = true)) -> (ms' = (VMS (p + 1%Z)%Z s m)))).
-
-Parameter beq: (func Z (func Z bool)).
-
-Axiom beq_def : forall (x:Z) (y:Z), ((infix_at (infix_at beq x)
-  y) = true) <-> (x = y).
-
-Axiom beq_lemma : forall (x:Z) (y:Z), ((infix_at (infix_at beq x)
-  y) = true) <-> (x = y).
-
-Parameter bne: (func Z (func Z bool)).
-
-Axiom bne_def : forall (x:Z) (y:Z), ((infix_at (infix_at bne x)
-  y) = true) <-> ~ (x = y).
-
-Axiom bne_lemma : forall (x:Z) (y:Z), ((infix_at (infix_at bne x)
-  y) = true) <-> ~ (x = y).
-
-Parameter ble: (func Z (func Z bool)).
-
-Axiom ble_def : forall (x:Z) (y:Z), ((infix_at (infix_at ble x)
-  y) = true) <-> (x <= y)%Z.
-
-Axiom ble_lemma : forall (x:Z) (y:Z), ((infix_at (infix_at ble x)
-  y) = true) <-> (x <= y)%Z.
-
-Parameter bgt: (func Z (func Z bool)).
-
-Axiom bgt_def : forall (x:Z) (y:Z), ((infix_at (infix_at bgt x)
-  y) = true) <-> (y < x)%Z.
-
-Axiom bgt_lemma : forall (x:Z) (y:Z), ((infix_at (infix_at bgt x)
-  y) = true) <-> (y < x)%Z.
-
-Parameter isetvar_pre: forall {a:Type} {a_WT:WhyType a}, (func a (func Z
-  (func machine_state bool))).
-
-Axiom isetvar_pre_def : forall {a:Type} {a_WT:WhyType a}, forall (x:a) (p:Z)
-  (ms:machine_state), ((infix_at (infix_at (infix_at (isetvar_pre : (func a
-  (func Z (func machine_state bool)))) x) p) ms) = true) <-> exists n:Z,
-  exists s:(list Z), exists m:(map id Z), (ms = (VMS p
-  (Init.Datatypes.cons n s) m)).
-
-Axiom isetvar_pre_lemma : forall {a:Type} {a_WT:WhyType a}, forall (x:a)
-  (p:Z) (ms:machine_state),
-  ((infix_at (infix_at (infix_at (isetvar_pre : (func a (func Z (func
-  machine_state bool)))) x) p) ms) = true) <-> exists n:Z, exists s:(list Z),
-  exists m:(map id Z), (ms = (VMS p (Init.Datatypes.cons n s) m)).
-
-Parameter isetvar_post: forall {a:Type} {a_WT:WhyType a}, id -> (func a (func
-  Z (func machine_state (func machine_state bool)))).
-
-Axiom isetvar_post_def : forall {a:Type} {a_WT:WhyType a}, forall (x:id)
-  (a1:a) (p:Z) (ms:machine_state) (ms':machine_state),
-  ((infix_at (infix_at (infix_at (infix_at (isetvar_post x: (func a (func Z
-  (func machine_state (func machine_state bool))))) a1) p) ms)
-  ms') = true) <-> forall (s:(list Z)) (n:Z) (m:(map id Z)), (ms = (VMS p
-  (Init.Datatypes.cons n s) m)) -> (ms' = (VMS (p + 1%Z)%Z s (set m x n))).
-
-Axiom isetvar_post_lemma : forall {a:Type} {a_WT:WhyType a}, forall (x:id)
-  (a1:a) (p:Z) (ms:machine_state) (ms':machine_state),
-  ((infix_at (infix_at (infix_at (infix_at (isetvar_post x: (func a (func Z
-  (func machine_state (func machine_state bool))))) a1) p) ms)
-  ms') = true) <-> forall (s:(list Z)) (n:Z) (m:(map id Z)), (ms = (VMS p
-  (Init.Datatypes.cons n s) m)) -> (ms' = (VMS (p + 1%Z)%Z s (set m x n))).
-
-Parameter aexpr_post: forall {a:Type} {a_WT:WhyType a}, aexpr -> Z -> (func a
-  (func Z (func machine_state (func machine_state bool)))).
-
-Axiom aexpr_post_def : forall {a:Type} {a_WT:WhyType a}, forall (a1:aexpr)
-  (len:Z) (x:a) (p:Z) (ms:machine_state) (ms':machine_state),
-  ((infix_at (infix_at (infix_at (infix_at (aexpr_post a1 len: (func a (func
-  Z (func machine_state (func machine_state bool))))) x) p) ms)
-  ms') = true) <->
-  match ms with
-  | (VMS _ s m) => (ms' = (VMS (p + len)%Z (Init.Datatypes.cons (aeval m
-      a1) s) m))
-  end.
-
-Axiom aexpr_post_lemma : forall {a:Type} {a_WT:WhyType a}, forall (a1:aexpr)
-  (len:Z) (x:a) (p:Z) (ms:machine_state) (ms':machine_state),
-  ((infix_at (infix_at (infix_at (infix_at (aexpr_post a1 len: (func a (func
-  Z (func machine_state (func machine_state bool))))) x) p) ms)
-  ms') = true) <->
-  match ms with
-  | (VMS _ s m) => (ms' = (VMS (p + len)%Z (Init.Datatypes.cons (aeval m
-      a1) s) m))
-  end.
-
-Parameter bexpr_post: forall {a:Type} {a_WT:WhyType a}, bexpr -> bool -> Z ->
-  Z -> (func a (func Z (func machine_state (func machine_state bool)))).
-
-Axiom bexpr_post_def : forall {a:Type} {a_WT:WhyType a}, forall (b:bexpr)
-  (cond1:bool) (out_t:Z) (out_f:Z) (x:a) (p:Z) (ms:machine_state)
-  (ms':machine_state),
-  (((infix_at (infix_at (infix_at (infix_at (bexpr_post b cond1 out_t
-  out_f: (func a (func Z (func machine_state (func machine_state bool))))) x)
-  p) ms) ms') = true) ->
-  match ms with
-  | (VMS _ s m) => (((beval m b) = cond1) -> (ms' = (VMS (p + out_t)%Z s
-      m))) /\ ((~ ((beval m b) = cond1)) -> (ms' = (VMS (p + out_f)%Z s m)))
-  end) /\
-  (match ms with
-  | (VMS _ s m) => (((beval m b) = cond1) /\ (ms' = (VMS (p + out_t)%Z s
-      m))) \/ ((~ ((beval m b) = cond1)) /\ (ms' = (VMS (p + out_f)%Z s m)))
-  end -> ((infix_at (infix_at (infix_at (infix_at (bexpr_post b cond1 out_t
-  out_f: (func a (func Z (func machine_state (func machine_state bool))))) x)
-  p) ms) ms') = true)).
-
-Parameter exn_bool: forall {a:Type} {a_WT:WhyType a}, bexpr -> bool -> (func
-  a (func Z (func machine_state bool))).
-
-Axiom exn_bool_def : forall {a:Type} {a_WT:WhyType a}, forall (b1:bexpr)
-  (cond1:bool) (x:a) (p:Z) (ms:machine_state),
-  ((infix_at (infix_at (infix_at (exn_bool b1 cond1: (func a (func Z (func
-  machine_state bool)))) x) p) ms) = true) <->
-  match ms with
-  | (VMS _ _ m) => ((beval m b1) = cond1)
-  end.
-
-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.
-
-Axiom com_pre_lemma : 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.
-
-Axiom com_post_lemma : 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.
-
-Parameter exn_bool_old: forall {a:Type} {a_WT:WhyType a}, bexpr -> bool ->
-  (func (a* machine_state)%type (func Z (func machine_state bool))).
-
-Axiom exn_bool_old_def : forall {a:Type} {a_WT:WhyType a}, forall (b1:bexpr)
-  (cond1:bool) (x:(a* machine_state)%type) (p:Z) (ms:machine_state),
-  ((infix_at (infix_at (infix_at (exn_bool_old b1 cond1: (func (a*
-  machine_state)%type (func Z (func machine_state bool)))) x) p)
-  ms) = true) <->
-  match (snd x) with
-  | (VMS _ _ m) => ((beval m b1) = cond1)
-  end.
-
-Parameter loop_invariant: forall {a:Type} {a_WT:WhyType a}, com -> (func (a*
-  machine_state)%type (func Z (func machine_state bool))).
-
-Axiom loop_invariant_def : forall {a:Type} {a_WT:WhyType a}, forall (c:com)
-  (x:(a* machine_state)%type) (p:Z) (msi:machine_state),
-  ((infix_at (infix_at (infix_at (loop_invariant c: (func (a*
-  machine_state)%type (func Z (func machine_state bool)))) x) p)
-  msi) = true) <-> match ((snd x),
-  msi) with
-  | ((VMS _ s0 m0), (VMS pi si mi)) => (pi = p) /\ ((s0 = si) /\
-      exists mf:(map id Z), (ceval m0 c mf) /\ (ceval mi c mf))
-  end.
-
-Parameter loop_post: forall {a:Type} {a_WT:WhyType a}, com -> Z -> (func (a*
-  machine_state)%type (func Z (func machine_state bool))).
-
-Axiom loop_post_def : forall {a:Type} {a_WT:WhyType a}, forall (c:com)
-  (len:Z) (x:(a* machine_state)%type) (p:Z) (msf:machine_state),
-  ((infix_at (infix_at (infix_at (loop_post c len: (func (a*
-  machine_state)%type (func Z (func machine_state bool)))) x) p)
-  msf) = true) <-> match ((snd x),
-  msf) with
-  | ((VMS _ s0 m0), (VMS pf sf mf)) => (pf = (p + len)%Z) /\ ((s0 = sf) /\
-      (ceval m0 c mf))
-  end.
-
-Parameter loop_variant: forall {a:Type} {a_WT:WhyType a}, com -> bexpr ->
-  (func a (func Z (func machine_state (func machine_state bool)))).
-
-Axiom loop_variant_def : forall {a:Type} {a_WT:WhyType a}, forall (c:com)
-  (test:bexpr) (x:a) (p:Z) (msj:machine_state) (msi:machine_state),
-  ((infix_at (infix_at (infix_at (infix_at (loop_variant c test: (func a
-  (func Z (func machine_state (func machine_state bool))))) x) p) msj)
-  msi) = true) <-> match (msj,
-  msi) with
-  | ((VMS pj sj mj), (VMS pi si mi)) => (pj = pi) /\ ((sj = si) /\ ((ceval mi
-      c mj) /\ ((beval mi test) = true)))
-  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')))).
-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.
-
diff --git a/examples/in_progress/mini-compiler/compiler/compiler_Compile_com_WP_parameter_compile_com_natural_2.v b/examples/in_progress/mini-compiler/compiler/compiler_Compile_com_WP_parameter_compile_com_natural_2.v
deleted file mode 100644
index f790cd8d57c2f7c0388fb4c792b9693b46f7fb77..0000000000000000000000000000000000000000
--- a/examples/in_progress/mini-compiler/compiler/compiler_Compile_com_WP_parameter_compile_com_natural_2.v
+++ /dev/null
@@ -1,249 +0,0 @@
-(* 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.
-
diff --git a/examples/in_progress/mini-compiler/compiler/why3session.xml b/examples/in_progress/mini-compiler/compiler/why3session.xml
index c04d732bee5b27249cdcb9c3687abfe6256412b0..ca28117598bf4ebcedf36c98052f7c02fd9db0c1 100644
--- a/examples/in_progress/mini-compiler/compiler/why3session.xml
+++ b/examples/in_progress/mini-compiler/compiler/why3session.xml
@@ -2,22 +2,20 @@
 <!DOCTYPE why3session PUBLIC "-//Why3//proof session v5//EN"
 "http://why3.lri.fr/why3session.dtd">
 <why3session shape_version="4">
-<prover id="0" name="Spass" version="3.5" timelimit="5" memlimit="1000"/>
+<prover id="0" name="CVC3" version="2.4.1" timelimit="5" memlimit="1000"/>
 <prover id="1" name="Coq" version="8.4pl2" timelimit="5" memlimit="1000"/>
 <prover id="2" name="Z3" version="4.3.1" timelimit="5" memlimit="1000"/>
-<prover id="3" name="Alt-Ergo" version="0.95.2" timelimit="5" memlimit="1000"/>
+<prover id="3" name="Eprover" version="1.8-001" timelimit="30" memlimit="1000"/>
+<prover id="4" name="Alt-Ergo" version="0.95.2" timelimit="5" memlimit="1000"/>
 <file name="../compiler.mlw" expanded="true">
-<theory name="Compile_aexpr" sum="4f7c3c3132f64ea43d0431210e906a5f">
- <goal name="aexpr_post_lemma">
- <proof prover="3"><result status="valid" time="0.06"/></proof>
- </goal>
+<theory name="Compile_aexpr" sum="5dad7bcfe154db92efe2a9d47ec3bcf4">
  <goal name="WP_parameter compile_aexpr" expl="VC for compile_aexpr">
  <transf name="split_goal_wp">
   <goal name="WP_parameter compile_aexpr.1" expl="1. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.2" expl="2. precondition">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.3" expl="3. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -26,10 +24,10 @@
     <goal name="WP_parameter compile_aexpr.3.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_aexpr.3.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.05"/></proof>
+     <proof prover="4"><result status="valid" time="0.05"/></proof>
      </goal>
      <goal name="WP_parameter compile_aexpr.3.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.07"/></proof>
+     <proof prover="4"><result status="valid" time="0.07"/></proof>
      </goal>
     </transf>
     </goal>
@@ -38,19 +36,19 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_aexpr.4" expl="4. postcondition">
-  <proof prover="3"><result status="valid" time="0.08"/></proof>
+  <proof prover="4"><result status="valid" time="0.08"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.5" expl="5. postcondition">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.6" expl="6. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.7" expl="7. precondition">
-  <proof prover="3"><result status="valid" time="0.04"/></proof>
+  <proof prover="4"><result status="valid" time="0.04"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.8" expl="8. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.9" expl="9. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -59,10 +57,10 @@
     <goal name="WP_parameter compile_aexpr.9.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_aexpr.9.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.06"/></proof>
+     <proof prover="4"><result status="valid" time="0.06"/></proof>
      </goal>
      <goal name="WP_parameter compile_aexpr.9.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.07"/></proof>
+     <proof prover="4"><result status="valid" time="0.07"/></proof>
      </goal>
     </transf>
     </goal>
@@ -71,37 +69,37 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_aexpr.10" expl="10. postcondition">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.11" expl="11. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.12" expl="12. postcondition">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.13" expl="13. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.14" expl="14. variant decrease">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.15" expl="15. precondition">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.16" expl="16. variant decrease">
-  <proof prover="3"><result status="valid" time="0.08"/></proof>
+  <proof prover="4"><result status="valid" time="0.08"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.17" expl="17. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.18" expl="18. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.19" expl="19. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.20" expl="20. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.21" expl="21. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -110,16 +108,16 @@
     <goal name="WP_parameter compile_aexpr.21.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_aexpr.21.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.08"/></proof>
+     <proof prover="4"><result status="valid" time="0.08"/></proof>
      </goal>
      <goal name="WP_parameter compile_aexpr.21.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.09"/></proof>
+     <proof prover="4"><result status="valid" time="0.09"/></proof>
      </goal>
      <goal name="WP_parameter compile_aexpr.21.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.09"/></proof>
+     <proof prover="4"><result status="valid" time="0.19"/></proof>
      </goal>
      <goal name="WP_parameter compile_aexpr.21.1.1.4" expl="4.">
-     <proof prover="3"><result status="valid" time="0.19"/></proof>
+     <proof prover="0"><result status="valid" time="0.12"/></proof>
      </goal>
     </transf>
     </goal>
@@ -128,37 +126,37 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_aexpr.22" expl="22. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.23" expl="23. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.24" expl="24. postcondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.25" expl="25. precondition">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.26" expl="26. variant decrease">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.27" expl="27. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.28" expl="28. variant decrease">
-  <proof prover="3"><result status="valid" time="0.08"/></proof>
+  <proof prover="4"><result status="valid" time="0.08"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.29" expl="29. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.30" expl="30. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.31" expl="31. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.32" expl="32. precondition">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.33" expl="33. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -167,16 +165,16 @@
     <goal name="WP_parameter compile_aexpr.33.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_aexpr.33.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.08"/></proof>
+     <proof prover="4"><result status="valid" time="0.08"/></proof>
      </goal>
      <goal name="WP_parameter compile_aexpr.33.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.09"/></proof>
+     <proof prover="4"><result status="valid" time="0.09"/></proof>
      </goal>
      <goal name="WP_parameter compile_aexpr.33.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.10"/></proof>
+     <proof prover="4"><result status="valid" time="0.09"/></proof>
      </goal>
      <goal name="WP_parameter compile_aexpr.33.1.1.4" expl="4.">
-     <proof prover="3"><result status="valid" time="0.09"/></proof>
+     <proof prover="0"><result status="valid" time="0.11"/></proof>
      </goal>
     </transf>
     </goal>
@@ -185,37 +183,37 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_aexpr.34" expl="34. postcondition">
-  <proof prover="3"><result status="valid" time="0.09"/></proof>
+  <proof prover="4"><result status="valid" time="0.09"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.35" expl="35. postcondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.36" expl="36. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.37" expl="37. precondition">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.38" expl="38. variant decrease">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.39" expl="39. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.40" expl="40. variant decrease">
-  <proof prover="3"><result status="valid" time="0.08"/></proof>
+  <proof prover="4"><result status="valid" time="0.08"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.41" expl="41. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.42" expl="42. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.43" expl="43. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.44" expl="44. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.45" expl="45. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -224,16 +222,16 @@
     <goal name="WP_parameter compile_aexpr.45.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_aexpr.45.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.08"/></proof>
+     <proof prover="4"><result status="valid" time="0.08"/></proof>
      </goal>
      <goal name="WP_parameter compile_aexpr.45.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.07"/></proof>
+     <proof prover="4"><result status="valid" time="0.07"/></proof>
      </goal>
      <goal name="WP_parameter compile_aexpr.45.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.10"/></proof>
+     <proof prover="4"><result status="valid" time="0.11"/></proof>
      </goal>
      <goal name="WP_parameter compile_aexpr.45.1.1.4" expl="4.">
-     <proof prover="3"><result status="valid" time="0.28"/></proof>
+     <proof prover="0"><result status="valid" time="0.10"/></proof>
      </goal>
     </transf>
     </goal>
@@ -242,35 +240,35 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_aexpr.46" expl="46. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.47" expl="47. postcondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr.48" expl="48. postcondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
  </transf>
  </goal>
  <goal name="WP_parameter compile_aexpr_natural" expl="VC for compile_aexpr_natural">
  <transf name="split_goal_wp">
   <goal name="WP_parameter compile_aexpr_natural.1" expl="1. assertion">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_aexpr_natural.2" expl="2. postcondition">
-  <proof prover="3"><result status="valid" time="0.12"/></proof>
+  <proof prover="4"><result status="valid" time="0.12"/></proof>
   </goal>
  </transf>
  </goal>
 </theory>
-<theory name="Compile_bexpr" sum="f981f82e61754462b4458823fb64fe7b">
+<theory name="Compile_bexpr" sum="e7457faac24e873b6f3ce07cc1d7426b">
  <goal name="WP_parameter compile_bexpr" expl="VC for compile_bexpr">
  <transf name="split_goal_wp">
   <goal name="WP_parameter compile_bexpr.1" expl="1. precondition">
-  <proof prover="3"><result status="valid" time="0.08"/></proof>
+  <proof prover="4"><result status="valid" time="0.08"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.2" expl="2. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.3" expl="3. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -279,13 +277,13 @@
     <goal name="WP_parameter compile_bexpr.3.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_bexpr.3.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.06"/></proof>
+     <proof prover="4"><result status="valid" time="0.06"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.3.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.06"/></proof>
+     <proof prover="4"><result status="valid" time="0.07"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.3.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.07"/></proof>
+     <proof prover="4"><result status="valid" time="0.06"/></proof>
      </goal>
     </transf>
     </goal>
@@ -294,19 +292,19 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_bexpr.4" expl="4. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.5" expl="5. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.6" expl="6. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.7" expl="7. precondition">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.8" expl="8. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.9" expl="9. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -315,13 +313,13 @@
     <goal name="WP_parameter compile_bexpr.9.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_bexpr.9.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.06"/></proof>
+     <proof prover="4"><result status="valid" time="0.06"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.9.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.07"/></proof>
+     <proof prover="4"><result status="valid" time="0.06"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.9.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.06"/></proof>
+     <proof prover="4"><result status="valid" time="0.07"/></proof>
      </goal>
     </transf>
     </goal>
@@ -330,19 +328,19 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_bexpr.10" expl="10. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.11" expl="11. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.12" expl="12. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.13" expl="13. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.14" expl="14. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.15" expl="15. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -351,13 +349,13 @@
     <goal name="WP_parameter compile_bexpr.15.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_bexpr.15.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.05"/></proof>
+     <proof prover="4"><result status="valid" time="0.05"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.15.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.08"/></proof>
+     <proof prover="4"><result status="valid" time="0.06"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.15.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.06"/></proof>
+     <proof prover="4"><result status="valid" time="0.08"/></proof>
      </goal>
     </transf>
     </goal>
@@ -366,19 +364,19 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_bexpr.16" expl="16. postcondition">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.17" expl="17. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.18" expl="18. postcondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.19" expl="19. precondition">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.20" expl="20. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.21" expl="21. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -387,13 +385,13 @@
     <goal name="WP_parameter compile_bexpr.21.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_bexpr.21.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.05"/></proof>
+     <proof prover="4"><result status="valid" time="0.05"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.21.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.07"/></proof>
+     <proof prover="4"><result status="valid" time="0.06"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.21.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.06"/></proof>
+     <proof prover="4"><result status="valid" time="0.07"/></proof>
      </goal>
     </transf>
     </goal>
@@ -402,22 +400,22 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_bexpr.22" expl="22. postcondition">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.23" expl="23. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.24" expl="24. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.25" expl="25. variant decrease">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.26" expl="26. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.27" expl="27. precondition">
-  <proof prover="3"><result status="valid" time="0.04"/></proof>
+  <proof prover="4"><result status="valid" time="0.04"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.28" expl="28. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -426,13 +424,13 @@
     <goal name="WP_parameter compile_bexpr.28.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_bexpr.28.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.06"/></proof>
+     <proof prover="4"><result status="valid" time="0.06"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.28.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.10"/></proof>
+     <proof prover="4"><result status="valid" time="0.10"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.28.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.10"/></proof>
+     <proof prover="4"><result status="valid" time="0.10"/></proof>
      </goal>
     </transf>
     </goal>
@@ -441,34 +439,34 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_bexpr.29" expl="29. postcondition">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.30" expl="30. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.31" expl="31. postcondition">
-  <proof prover="3"><result status="valid" time="0.04"/></proof>
+  <proof prover="4"><result status="valid" time="0.04"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.32" expl="32. variant decrease">
-  <proof prover="3"><result status="valid" time="0.08"/></proof>
+  <proof prover="4"><result status="valid" time="0.08"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.33" expl="33. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.34" expl="34. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.35" expl="35. variant decrease">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.36" expl="36. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.37" expl="37. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.38" expl="38. precondition">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.39" expl="39. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -477,22 +475,22 @@
     <goal name="WP_parameter compile_bexpr.39.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_bexpr.39.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.23"/></proof>
+     <proof prover="4"><result status="valid" time="0.10"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.39.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.20"/></proof>
+     <proof prover="4"><result status="valid" time="0.23"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.39.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.27"/></proof>
+     <proof prover="4"><result status="valid" time="0.28"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.39.1.1.4" expl="4.">
-     <proof prover="3"><result status="valid" time="0.12"/></proof>
+     <proof prover="4"><result status="valid" time="0.12"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.39.1.1.5" expl="5.">
-     <proof prover="3"><result status="valid" time="0.45"/></proof>
+     <proof prover="4"><result status="valid" time="0.41"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.39.1.1.6" expl="6.">
-     <proof prover="3"><result status="valid" time="0.45"/></proof>
+     <proof prover="4"><result status="valid" time="0.47"/></proof>
      </goal>
     </transf>
     </goal>
@@ -501,31 +499,31 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_bexpr.40" expl="40. postcondition">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.41" expl="41. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.42" expl="42. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.43" expl="43. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.44" expl="44. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.45" expl="45. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.46" expl="46. precondition">
-  <proof prover="3"><result status="valid" time="0.09"/></proof>
+  <proof prover="4"><result status="valid" time="0.09"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.47" expl="47. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.48" expl="48. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.49" expl="49. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -534,19 +532,19 @@
     <goal name="WP_parameter compile_bexpr.49.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_bexpr.49.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.09"/></proof>
+     <proof prover="4"><result status="valid" time="0.09"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.49.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.08"/></proof>
+     <proof prover="4"><result status="valid" time="0.08"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.49.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.21"/></proof>
+     <proof prover="4"><result status="valid" time="0.09"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.49.1.1.4" expl="4.">
-     <proof prover="3"><result status="valid" time="0.27"/></proof>
+     <proof prover="0"><result status="valid" time="0.12"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.49.1.1.5" expl="5.">
-     <proof prover="3"><result status="valid" time="0.28"/></proof>
+     <proof prover="0"><result status="valid" time="0.13"/></proof>
      </goal>
     </transf>
     </goal>
@@ -555,31 +553,31 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_bexpr.50" expl="50. postcondition">
-  <proof prover="3"><result status="valid" time="0.10"/></proof>
+  <proof prover="4"><result status="valid" time="0.10"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.51" expl="51. postcondition">
-  <proof prover="3"><result status="valid" time="0.09"/></proof>
+  <proof prover="4"><result status="valid" time="0.09"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.52" expl="52. postcondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.53" expl="53. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.54" expl="54. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.55" expl="55. precondition">
-  <proof prover="3"><result status="valid" time="0.08"/></proof>
+  <proof prover="4"><result status="valid" time="0.08"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.56" expl="56. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.57" expl="57. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.58" expl="58. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.59" expl="59. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -588,19 +586,19 @@
     <goal name="WP_parameter compile_bexpr.59.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_bexpr.59.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.08"/></proof>
+     <proof prover="4"><result status="valid" time="0.08"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.59.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.09"/></proof>
+     <proof prover="4"><result status="valid" time="0.09"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.59.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.11"/></proof>
+     <proof prover="4"><result status="valid" time="0.08"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.59.1.1.4" expl="4.">
-     <proof prover="3"><result status="valid" time="0.25"/></proof>
+     <proof prover="0"><result status="valid" time="0.11"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.59.1.1.5" expl="5.">
-     <proof prover="3"><result status="valid" time="0.24"/></proof>
+     <proof prover="0"><result status="valid" time="0.10"/></proof>
      </goal>
     </transf>
     </goal>
@@ -609,31 +607,31 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_bexpr.60" expl="60. postcondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.61" expl="61. postcondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.62" expl="62. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.63" expl="63. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.64" expl="64. precondition">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.65" expl="65. precondition">
-  <proof prover="3"><result status="valid" time="0.08"/></proof>
+  <proof prover="4"><result status="valid" time="0.08"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.66" expl="66. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.67" expl="67. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.68" expl="68. precondition">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.69" expl="69. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -642,19 +640,19 @@
     <goal name="WP_parameter compile_bexpr.69.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_bexpr.69.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.08"/></proof>
+     <proof prover="4"><result status="valid" time="0.11"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.69.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.09"/></proof>
+     <proof prover="4"><result status="valid" time="0.09"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.69.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.12"/></proof>
+     <proof prover="4"><result status="valid" time="0.08"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.69.1.1.4" expl="4.">
-     <proof prover="3"><result status="valid" time="0.26"/></proof>
+     <proof prover="0"><result status="valid" time="0.13"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.69.1.1.5" expl="5.">
-     <proof prover="3"><result status="valid" time="0.24"/></proof>
+     <proof prover="0"><result status="valid" time="0.17"/></proof>
      </goal>
     </transf>
     </goal>
@@ -663,31 +661,31 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_bexpr.70" expl="70. postcondition">
-  <proof prover="3"><result status="valid" time="0.09"/></proof>
+  <proof prover="4"><result status="valid" time="0.09"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.71" expl="71. postcondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.72" expl="72. postcondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.73" expl="73. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.74" expl="74. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.75" expl="75. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.76" expl="76. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.77" expl="77. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.78" expl="78. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.79" expl="79. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -696,19 +694,19 @@
     <goal name="WP_parameter compile_bexpr.79.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_bexpr.79.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.07"/></proof>
+     <proof prover="4"><result status="valid" time="0.09"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.79.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.07"/></proof>
+     <proof prover="4"><result status="valid" time="0.07"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.79.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.23"/></proof>
+     <proof prover="4"><result status="valid" time="0.07"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.79.1.1.4" expl="4.">
-     <proof prover="3"><result status="valid" time="0.27"/></proof>
+     <proof prover="0"><result status="valid" time="0.12"/></proof>
      </goal>
      <goal name="WP_parameter compile_bexpr.79.1.1.5" expl="5.">
-     <proof prover="3"><result status="valid" time="0.34"/></proof>
+     <proof prover="0"><result status="valid" time="0.11"/></proof>
      </goal>
     </transf>
     </goal>
@@ -717,34 +715,28 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_bexpr.80" expl="80. postcondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.81" expl="81. postcondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_bexpr.82" expl="82. postcondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
  </transf>
  </goal>
  <goal name="WP_parameter compile_bexpr_natural" expl="VC for compile_bexpr_natural">
- <proof prover="3"><result status="valid" time="2.74"/></proof>
+ <proof prover="4"><result status="valid" time="3.51"/></proof>
  </goal>
 </theory>
-<theory name="Compile_com" sum="8fdd042031fb4c236cbb7a8b7269e53c" expanded="true">
- <goal name="com_pre_lemma">
- <proof prover="3"><result status="valid" time="0.09"/></proof>
- </goal>
- <goal name="com_post_lemma">
- <proof prover="3"><result status="valid" time="0.12"/></proof>
- </goal>
+<theory name="Compile_com" sum="2dfe5b2cfda1fd51d4e97bbb62be8b66">
  <goal name="WP_parameter compile_com" expl="VC for compile_com">
  <transf name="split_goal_wp">
   <goal name="WP_parameter compile_com.1" expl="1. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.2" expl="2. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.3" expl="3. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -753,16 +745,16 @@
     <goal name="WP_parameter compile_com.3.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_com.3.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.08"/></proof>
+     <proof prover="4"><result status="valid" time="0.10"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.3.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.07"/></proof>
+     <proof prover="4"><result status="valid" time="0.08"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.3.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.10"/></proof>
+     <proof prover="4"><result status="valid" time="0.07"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.3.1.1.4" expl="4.">
-     <proof prover="3"><result status="valid" time="0.08"/></proof>
+     <proof prover="4"><result status="valid" time="0.08"/></proof>
      </goal>
     </transf>
     </goal>
@@ -771,25 +763,25 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_com.4" expl="4. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.5" expl="5. postcondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.6" expl="6. postcondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.7" expl="7. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.8" expl="8. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.9" expl="9. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.10" expl="10. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.11" expl="11. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -798,19 +790,19 @@
     <goal name="WP_parameter compile_com.11.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_com.11.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.08"/></proof>
+     <proof prover="4"><result status="valid" time="0.08"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.11.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.11"/></proof>
+     <proof prover="4"><result status="valid" time="0.11"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.11.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.11"/></proof>
+     <proof prover="4"><result status="valid" time="0.12"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.11.1.1.4" expl="4.">
-     <proof prover="3"><result status="valid" time="0.12"/></proof>
+     <proof prover="4"><result status="valid" time="0.11"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.11.1.1.5" expl="5.">
-     <proof prover="3"><result status="valid" time="0.54"/></proof>
+     <proof prover="4"><result status="valid" time="0.54"/></proof>
      </goal>
     </transf>
     </goal>
@@ -819,31 +811,31 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_com.12" expl="12. postcondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.13" expl="13. postcondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.14" expl="14. postcondition">
-  <proof prover="3"><result status="valid" time="0.08"/></proof>
+  <proof prover="4"><result status="valid" time="0.08"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.15" expl="15. variant decrease">
-  <proof prover="3"><result status="valid" time="0.08"/></proof>
+  <proof prover="4"><result status="valid" time="0.08"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.16" expl="16. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.17" expl="17. variant decrease">
-  <proof prover="3"><result status="valid" time="0.08"/></proof>
+  <proof prover="4"><result status="valid" time="0.08"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.18" expl="18. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.19" expl="19. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.20" expl="20. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.21" expl="21. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -852,25 +844,25 @@
     <goal name="WP_parameter compile_com.21.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_com.21.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.08"/></proof>
+     <proof prover="4"><result status="valid" time="0.08"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.21.1.1.2" expl="2.">
-     <proof prover="2"><result status="valid" time="0.05"/></proof>
+     <proof prover="2"><result status="valid" time="0.06"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.21.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.08"/></proof>
+     <proof prover="4"><result status="valid" time="0.08"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.21.1.1.4" expl="4.">
-     <proof prover="2"><result status="valid" time="0.06"/></proof>
+     <proof prover="2"><result status="valid" time="0.05"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.21.1.1.5" expl="5.">
-     <proof prover="3"><result status="valid" time="0.09"/></proof>
+     <proof prover="4"><result status="valid" time="0.13"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.21.1.1.6" expl="6.">
-     <proof prover="3"><result status="valid" time="0.13"/></proof>
+     <proof prover="4"><result status="valid" time="0.09"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.21.1.1.7" expl="7.">
-     <proof prover="3"><result status="valid" time="0.09"/></proof>
+     <proof prover="4"><result status="valid" time="0.09"/></proof>
      </goal>
     </transf>
     </goal>
@@ -879,49 +871,49 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_com.22" expl="22. postcondition">
-  <proof prover="3"><result status="valid" time="0.05"/></proof>
+  <proof prover="4"><result status="valid" time="0.05"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.23" expl="23. postcondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.24" expl="24. postcondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.25" expl="25. variant decrease">
-  <proof prover="3"><result status="valid" time="0.09"/></proof>
+  <proof prover="4"><result status="valid" time="0.09"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.26" expl="26. variant decrease">
-  <proof prover="3"><result status="valid" time="0.09"/></proof>
+  <proof prover="4"><result status="valid" time="0.09"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.27" expl="27. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.28" expl="28. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.29" expl="29. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.30" expl="30. precondition">
-  <proof prover="3"><result status="valid" time="0.08"/></proof>
+  <proof prover="4"><result status="valid" time="0.08"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.31" expl="31. precondition">
-  <proof prover="3"><result status="valid" time="0.07"/></proof>
+  <proof prover="4"><result status="valid" time="0.07"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.32" expl="32. precondition">
-  <proof prover="3"><result status="valid" time="0.08"/></proof>
+  <proof prover="4"><result status="valid" time="0.08"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.33" expl="33. precondition">
-  <proof prover="3"><result status="valid" time="0.09"/></proof>
+  <proof prover="4"><result status="valid" time="0.09"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.34" expl="34. precondition">
-  <proof prover="3"><result status="valid" time="0.09"/></proof>
+  <proof prover="4"><result status="valid" time="0.09"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.35" expl="35. precondition">
-  <proof prover="3"><result status="valid" time="0.09"/></proof>
+  <proof prover="4"><result status="valid" time="0.09"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.36" expl="36. precondition">
-  <proof prover="3"><result status="valid" time="0.06"/></proof>
+  <proof prover="4"><result status="valid" time="0.06"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.37" expl="37. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -930,64 +922,64 @@
     <goal name="WP_parameter compile_com.37.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_com.37.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.25"/></proof>
+     <proof prover="4"><result status="valid" time="0.25"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.39"/></proof>
+     <proof prover="4"><result status="valid" time="0.50"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.31"/></proof>
+     <proof prover="4"><result status="valid" time="0.44"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.4" expl="4.">
-     <proof prover="3"><result status="valid" time="0.30"/></proof>
+     <proof prover="4"><result status="valid" time="0.47"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.5" expl="5.">
-     <proof prover="3"><result status="valid" time="0.25"/></proof>
+     <proof prover="4"><result status="valid" time="0.52"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.6" expl="6.">
-     <proof prover="3"><result status="valid" time="1.07"/></proof>
+     <proof prover="4"><result status="valid" time="1.18"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.7" expl="7.">
-     <proof prover="3"><result status="valid" time="0.28"/></proof>
+     <proof prover="4"><result status="valid" time="0.62"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.8" expl="8.">
-     <proof prover="3"><result status="valid" time="0.37"/></proof>
+     <proof prover="4"><result status="valid" time="0.51"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.9" expl="9.">
-     <proof prover="3"><result status="valid" time="0.34"/></proof>
+     <proof prover="4"><result status="valid" time="0.47"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.10" expl="10.">
-     <proof prover="3"><result status="valid" time="0.38"/></proof>
+     <proof prover="4"><result status="valid" time="0.41"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.11" expl="11.">
-     <proof prover="3"><result status="valid" time="0.98"/></proof>
+     <proof prover="4"><result status="valid" time="1.07"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.12" expl="12.">
-     <proof prover="2"><result status="valid" time="0.25"/></proof>
+     <proof prover="2"><result status="valid" time="0.39"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.13" expl="13.">
-     <proof prover="3"><result status="valid" time="0.38"/></proof>
+     <proof prover="4"><result status="valid" time="0.63"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.14" expl="14.">
-     <proof prover="3"><result status="valid" time="0.45"/></proof>
+     <proof prover="4"><result status="valid" time="0.63"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.15" expl="15.">
-     <proof prover="3"><result status="valid" time="0.40"/></proof>
+     <proof prover="4"><result status="valid" time="0.53"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.16" expl="16.">
-     <proof prover="3"><result status="valid" time="0.41"/></proof>
+     <proof prover="4"><result status="valid" time="0.38"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.17" expl="17.">
-     <proof prover="3"><result status="valid" time="0.30"/></proof>
+     <proof prover="4"><result status="valid" time="0.29"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.18" expl="18.">
-     <proof prover="3"><result status="valid" time="0.28"/></proof>
+     <proof prover="4"><result status="valid" time="0.38"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.19" expl="19.">
-     <proof prover="3"><result status="valid" time="0.35"/></proof>
+     <proof prover="4"><result status="valid" time="0.45"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.37.1.1.20" expl="20.">
-     <proof prover="3"><result status="valid" time="0.49"/></proof>
+     <proof prover="4"><result status="valid" time="0.40"/></proof>
      </goal>
     </transf>
     </goal>
@@ -996,37 +988,37 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_com.38" expl="38. postcondition">
-  <proof prover="3"><result status="valid" time="0.12"/></proof>
+  <proof prover="4"><result status="valid" time="0.12"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.39" expl="39. postcondition">
-  <proof prover="3"><result status="valid" time="0.09"/></proof>
+  <proof prover="4"><result status="valid" time="0.09"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.40" expl="40. postcondition">
-  <proof prover="3"><result status="valid" time="0.09"/></proof>
+  <proof prover="4"><result status="valid" time="0.09"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.41" expl="41. variant decrease">
-  <proof prover="3"><result status="valid" time="0.09"/></proof>
+  <proof prover="4"><result status="valid" time="0.09"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.42" expl="42. precondition">
-  <proof prover="3"><result status="valid" time="0.09"/></proof>
+  <proof prover="4"><result status="valid" time="0.09"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.43" expl="43. precondition">
-  <proof prover="3"><result status="valid" time="0.08"/></proof>
+  <proof prover="4"><result status="valid" time="0.08"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.44" expl="44. precondition">
-  <proof prover="3"><result status="valid" time="0.08"/></proof>
+  <proof prover="4"><result status="valid" time="0.08"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.45" expl="45. precondition">
-  <proof prover="3"><result status="valid" time="0.09"/></proof>
+  <proof prover="4"><result status="valid" time="0.09"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.46" expl="46. precondition">
-  <proof prover="3"><result status="valid" time="0.09"/></proof>
+  <proof prover="4"><result status="valid" time="0.09"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.47" expl="47. precondition">
-  <proof prover="3"><result status="valid" time="0.08"/></proof>
+  <proof prover="4"><result status="valid" time="0.08"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.48" expl="48. precondition">
-  <proof prover="3"><result status="valid" time="0.09"/></proof>
+  <proof prover="4"><result status="valid" time="0.09"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.49" expl="49. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -1035,22 +1027,787 @@
     <goal name="WP_parameter compile_com.49.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_com.49.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.43"/></proof>
+     <proof prover="4"><result status="valid" time="0.47"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.49.1.1.2" expl="2.">
-     <proof prover="2"><result status="valid" time="0.14"/></proof>
+     <proof prover="2"><result status="valid" time="0.29"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.49.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.34"/></proof>
+     <proof prover="4"><result status="valid" time="0.34"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.49.1.1.4" expl="4.">
      <proof prover="2"><result status="valid" time="0.12"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.49.1.1.5" expl="5.">
-     <proof prover="3"><result status="valid" time="0.44"/></proof>
+     <proof prover="4"><result status="valid" time="0.44"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.49.1.1.6" expl="6.">
-     <proof prover="2"><result status="valid" time="2.12"/></proof>
+     <metas>
+      <ts_pos name="real" arity="0" id="2"
+       ip_theory="BuiltIn">
+       <ip_library name="why3"/>
+       <ip_library name="BuiltIn"/>
+       <ip_qualid name="real"/>
+      </ts_pos>
+      <ts_pos name="pred" arity="1" id="8"
+       ip_theory="HighOrd">
+       <ip_library name="why3"/>
+       <ip_library name="HighOrd"/>
+       <ip_qualid name="pred"/>
+      </ts_pos>
+      <ts_pos name="&apos;mark" arity="0" id="54"
+       ip_theory="Mark">
+       <ip_library name="why3"/>
+       <ip_library name="Mark"/>
+       <ip_qualid name="&apos;mark"/>
+      </ts_pos>
+      <ts_pos name="state" arity="0" id="4946"
+       ip_theory="State">
+       <ip_library name="state"/>
+       <ip_qualid name="state"/>
+      </ts_pos>
+      <ts_pos name="pos" arity="0" id="5227"
+       ip_theory="Vm">
+       <ip_library name="vm"/>
+       <ip_qualid name="pos"/>
+      </ts_pos>
+      <ts_pos name="stack" arity="0" id="5228"
+       ip_theory="Vm">
+       <ip_library name="vm"/>
+       <ip_qualid name="stack"/>
+      </ts_pos>
+      <ts_pos name="code" arity="0" id="5244"
+       ip_theory="Vm">
+       <ip_library name="vm"/>
+       <ip_qualid name="code"/>
+      </ts_pos>
+      <ts_pos name="pred" arity="0" id="5635"
+       ip_theory="Compiler_logic">
+       <ip_library name="logic"/>
+       <ip_qualid name="pred"/>
+      </ts_pos>
+      <ts_pos name="post" arity="0" id="5636"
+       ip_theory="Compiler_logic">
+       <ip_library name="logic"/>
+       <ip_qualid name="post"/>
+      </ts_pos>
+      <ts_pos name="binop" arity="0" id="6208"
+       ip_theory="VM_arith_instr_spec">
+       <ip_library name="specs"/>
+       <ip_qualid name="binop"/>
+      </ts_pos>
+      <ts_pos name="cond" arity="0" id="6495"
+       ip_theory="VM_bool_instr_spec">
+       <ip_library name="specs"/>
+       <ip_qualid name="cond"/>
+      </ts_pos>
+      <ls_pos name="infix =" id="10"
+       ip_theory="BuiltIn">
+       <ip_library name="why3"/>
+       <ip_library name="BuiltIn"/>
+       <ip_qualid name="infix ="/>
+      </ls_pos>
+      <ls_pos name="infix @" id="15"
+       ip_theory="HighOrd">
+       <ip_library name="why3"/>
+       <ip_library name="HighOrd"/>
+       <ip_qualid name="infix @"/>
+      </ls_pos>
+      <ls_pos name="one" id="689"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="one"/>
+      </ls_pos>
+      <ls_pos name="infix &lt;" id="690"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="infix &lt;"/>
+      </ls_pos>
+      <ls_pos name="infix &gt;" id="693"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="infix &gt;"/>
+      </ls_pos>
+      <ls_pos name="infix +" id="1859"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="infix +"/>
+      </ls_pos>
+      <ls_pos name="prefix -" id="1860"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="prefix -"/>
+      </ls_pos>
+      <ls_pos name="infix *" id="1861"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="infix *"/>
+      </ls_pos>
+      <ls_pos name="abs" id="1974"
+       ip_theory="Abs">
+       <ip_library name="int"/>
+       <ip_qualid name="abs"/>
+      </ls_pos>
+      <ls_pos name="div" id="2099"
+       ip_theory="EuclideanDivision">
+       <ip_library name="int"/>
+       <ip_qualid name="div"/>
+      </ls_pos>
+      <ls_pos name="mod" id="2102"
+       ip_theory="EuclideanDivision">
+       <ip_library name="int"/>
+       <ip_qualid name="mod"/>
+      </ls_pos>
+      <ls_pos name="orb" id="3648"
+       ip_theory="Bool">
+       <ip_library name="bool"/>
+       <ip_qualid name="orb"/>
+      </ls_pos>
+      <ls_pos name="xorb" id="3657"
+       ip_theory="Bool">
+       <ip_library name="bool"/>
+       <ip_qualid name="xorb"/>
+      </ls_pos>
+      <ls_pos name="implb" id="3671"
+       ip_theory="Bool">
+       <ip_library name="bool"/>
+       <ip_qualid name="implb"/>
+      </ls_pos>
+      <ls_pos name="get" id="4902"
+       ip_theory="State">
+       <ip_library name="state"/>
+       <ip_qualid name="get"/>
+      </ls_pos>
+      <ls_pos name="set" id="4903"
+       ip_theory="State">
+       <ip_library name="state"/>
+       <ip_qualid name="set"/>
+      </ls_pos>
+      <ls_pos name="const" id="4940"
+       ip_theory="State">
+       <ip_library name="state"/>
+       <ip_qualid name="const"/>
+      </ls_pos>
+      <ls_pos name="ihalt" id="5329"
+       ip_theory="Vm">
+       <ip_library name="vm"/>
+       <ip_qualid name="ihalt"/>
+      </ls_pos>
+      <ls_pos name="vm_terminates" id="5580"
+       ip_theory="Vm">
+       <ip_library name="vm"/>
+       <ip_qualid name="vm_terminates"/>
+      </ls_pos>
+      <ls_pos name="acc" id="5984"
+       ip_theory="Compiler_logic">
+       <ip_library name="logic"/>
+       <ip_qualid name="acc"/>
+      </ls_pos>
+      <ls_pos name="loop_preservation" id="5998"
+       ip_theory="Compiler_logic">
+       <ip_library name="logic"/>
+       <ip_qualid name="loop_preservation"/>
+      </ls_pos>
+      <ls_pos name="forget_old" id="6034"
+       ip_theory="Compiler_logic">
+       <ip_library name="logic"/>
+       <ip_qualid name="forget_old"/>
+      </ls_pos>
+      <ls_pos name="ibinop_post" id="6242"
+       ip_theory="VM_arith_instr_spec">
+       <ip_library name="specs"/>
+       <ip_qualid name="ibinop_post"/>
+      </ls_pos>
+      <ls_pos name="sub" id="6341"
+       ip_theory="VM_arith_instr_spec">
+       <ip_library name="specs"/>
+       <ip_qualid name="sub"/>
+      </ls_pos>
+      <ls_pos name="mul" id="6354"
+       ip_theory="VM_arith_instr_spec">
+       <ip_library name="specs"/>
+       <ip_qualid name="mul"/>
+      </ls_pos>
+      <ls_pos name="inil_post" id="6401"
+       ip_theory="VM_bool_instr_spec">
+       <ip_library name="specs"/>
+       <ip_qualid name="inil_post"/>
+      </ls_pos>
+      <ls_pos name="ibranch_post" id="6433"
+       ip_theory="VM_bool_instr_spec">
+       <ip_library name="specs"/>
+       <ip_qualid name="ibranch_post"/>
+      </ls_pos>
+      <ls_pos name="icjump_post" id="6496"
+       ip_theory="VM_bool_instr_spec">
+       <ip_library name="specs"/>
+       <ip_qualid name="icjump_post"/>
+      </ls_pos>
+      <ls_pos name="beq" id="6543"
+       ip_theory="VM_bool_instr_spec">
+       <ip_library name="specs"/>
+       <ip_qualid name="beq"/>
+      </ls_pos>
+      <ls_pos name="bne" id="6556"
+       ip_theory="VM_bool_instr_spec">
+       <ip_library name="specs"/>
+       <ip_qualid name="bne"/>
+      </ls_pos>
+      <ls_pos name="ble" id="6569"
+       ip_theory="VM_bool_instr_spec">
+       <ip_library name="specs"/>
+       <ip_qualid name="ble"/>
+      </ls_pos>
+      <ls_pos name="bgt" id="6582"
+       ip_theory="VM_bool_instr_spec">
+       <ip_library name="specs"/>
+       <ip_qualid name="bgt"/>
+      </ls_pos>
+      <ls_pos name="aexpr_post" id="6797"
+       ip_theory="Compile_aexpr">
+       <ip_qualid name="aexpr_post"/>
+      </ls_pos>
+      <ls_pos name="bexpr_post" id="8537"
+       ip_theory="Compile_bexpr">
+       <ip_qualid name="bexpr_post"/>
+      </ls_pos>
+      <ls_pos name="exn_bool" id="8596"
+       ip_theory="Compile_bexpr">
+       <ip_qualid name="exn_bool"/>
+      </ls_pos>
+      <ls_pos name="com_pre" id="10482"
+       ip_theory="Compile_com">
+       <ip_qualid name="com_pre"/>
+      </ls_pos>
+      <ls_pos name="com_post" id="10526"
+       ip_theory="Compile_com">
+       <ip_qualid name="com_post"/>
+      </ls_pos>
+      <ls_pos name="exn_bool_old" id="10605"
+       ip_theory="Compile_com">
+       <ip_qualid name="exn_bool_old"/>
+      </ls_pos>
+      <ls_pos name="loop_invariant" id="10644"
+       ip_theory="Compile_com">
+       <ip_qualid name="loop_invariant"/>
+      </ls_pos>
+      <ls_pos name="loop_post" id="10716"
+       ip_theory="Compile_com">
+       <ip_qualid name="loop_post"/>
+      </ls_pos>
+      <ls_pos name="loop_variant" id="10789"
+       ip_theory="Compile_com">
+       <ip_qualid name="loop_variant"/>
+      </ls_pos>
+      <pr_pos name="Assoc" id="1862"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="CommutativeGroup"/>
+       <ip_qualid name="Assoc"/>
+      </pr_pos>
+      <pr_pos name="Unit_def_l" id="1869"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="CommutativeGroup"/>
+       <ip_qualid name="Unit_def_l"/>
+      </pr_pos>
+      <pr_pos name="Unit_def_r" id="1872"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="CommutativeGroup"/>
+       <ip_qualid name="Unit_def_r"/>
+      </pr_pos>
+      <pr_pos name="Inv_def_l" id="1875"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="CommutativeGroup"/>
+       <ip_qualid name="Inv_def_l"/>
+      </pr_pos>
+      <pr_pos name="Inv_def_r" id="1878"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="CommutativeGroup"/>
+       <ip_qualid name="Inv_def_r"/>
+      </pr_pos>
+      <pr_pos name="Comm" id="1881"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="CommutativeGroup"/>
+       <ip_qualid name="Comm"/>
+       <ip_qualid name="Comm"/>
+      </pr_pos>
+      <pr_pos name="Assoc" id="1886"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="Assoc"/>
+       <ip_qualid name="Assoc"/>
+      </pr_pos>
+      <pr_pos name="Mul_distr_l" id="1893"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="Mul_distr_l"/>
+      </pr_pos>
+      <pr_pos name="Mul_distr_r" id="1900"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="Mul_distr_r"/>
+      </pr_pos>
+      <pr_pos name="Comm" id="1918"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="Comm"/>
+       <ip_qualid name="Comm"/>
+      </pr_pos>
+      <pr_pos name="Unitary" id="1923"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="Unitary"/>
+      </pr_pos>
+      <pr_pos name="NonTrivialRing" id="1926"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="NonTrivialRing"/>
+      </pr_pos>
+      <pr_pos name="Refl" id="1938"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="Refl"/>
+      </pr_pos>
+      <pr_pos name="Trans" id="1941"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="Trans"/>
+      </pr_pos>
+      <pr_pos name="Antisymm" id="1948"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="Antisymm"/>
+      </pr_pos>
+      <pr_pos name="Total" id="1953"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="Total"/>
+      </pr_pos>
+      <pr_pos name="ZeroLessOne" id="1958"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="ZeroLessOne"/>
+      </pr_pos>
+      <pr_pos name="CompatOrderAdd" id="1959"
+       ip_theory="Int">
+       <ip_library name="int"/>
+       <ip_qualid name="CompatOrderAdd"/>
+      </pr_pos>
+      <pr_pos name="Abs_le" id="1979"
+       ip_theory="Abs">
+       <ip_library name="int"/>
+       <ip_qualid name="Abs_le"/>
+      </pr_pos>
+      <pr_pos name="Abs_pos" id="1984"
+       ip_theory="Abs">
+       <ip_library name="int"/>
+       <ip_qualid name="Abs_pos"/>
+      </pr_pos>
+      <pr_pos name="Div_mod" id="2105"
+       ip_theory="EuclideanDivision">
+       <ip_library name="int"/>
+       <ip_qualid name="Div_mod"/>
+      </pr_pos>
+      <pr_pos name="Div_bound" id="2110"
+       ip_theory="EuclideanDivision">
+       <ip_library name="int"/>
+       <ip_qualid name="Div_bound"/>
+      </pr_pos>
+      <pr_pos name="Mod_bound" id="2115"
+       ip_theory="EuclideanDivision">
+       <ip_library name="int"/>
+       <ip_qualid name="Mod_bound"/>
+      </pr_pos>
+      <pr_pos name="Mod_1" id="2120"
+       ip_theory="EuclideanDivision">
+       <ip_library name="int"/>
+       <ip_qualid name="Mod_1"/>
+      </pr_pos>
+      <pr_pos name="Div_1" id="2123"
+       ip_theory="EuclideanDivision">
+       <ip_library name="int"/>
+       <ip_qualid name="Div_1"/>
+      </pr_pos>
+      <pr_pos name="Div_inf" id="2126"
+       ip_theory="EuclideanDivision">
+       <ip_library name="int"/>
+       <ip_qualid name="Div_inf"/>
+      </pr_pos>
+      <pr_pos name="Div_inf_neg" id="2131"
+       ip_theory="EuclideanDivision">
+       <ip_library name="int"/>
+       <ip_qualid name="Div_inf_neg"/>
+      </pr_pos>
+      <pr_pos name="Mod_0" id="2136"
+       ip_theory="EuclideanDivision">
+       <ip_library name="int"/>
+       <ip_qualid name="Mod_0"/>
+      </pr_pos>
+      <pr_pos name="Div_1_left" id="2139"
+       ip_theory="EuclideanDivision">
+       <ip_library name="int"/>
+       <ip_qualid name="Div_1_left"/>
+      </pr_pos>
+      <pr_pos name="Div_minus1_left" id="2142"
+       ip_theory="EuclideanDivision">
+       <ip_library name="int"/>
+       <ip_qualid name="Div_minus1_left"/>
+      </pr_pos>
+      <pr_pos name="Mod_1_left" id="2145"
+       ip_theory="EuclideanDivision">
+       <ip_library name="int"/>
+       <ip_qualid name="Mod_1_left"/>
+      </pr_pos>
+      <pr_pos name="Mod_minus1_left" id="2148"
+       ip_theory="EuclideanDivision">
+       <ip_library name="int"/>
+       <ip_qualid name="Mod_minus1_left"/>
+      </pr_pos>
+      <pr_pos name="Div_mult" id="2151"
+       ip_theory="EuclideanDivision">
+       <ip_library name="int"/>
+       <ip_qualid name="Div_mult"/>
+      </pr_pos>
+      <pr_pos name="Mod_mult" id="2158"
+       ip_theory="EuclideanDivision">
+       <ip_library name="int"/>
+       <ip_qualid name="Mod_mult"/>
+      </pr_pos>
+      <pr_pos name="Append_assoc" id="4292"
+       ip_theory="Append">
+       <ip_library name="list"/>
+       <ip_qualid name="Append_assoc"/>
+      </pr_pos>
+      <pr_pos name="ceval_deterministic_aux" id="5151"
+       ip_theory="Imp">
+       <ip_library name="imp"/>
+       <ip_qualid name="ceval_deterministic_aux"/>
+      </pr_pos>
+      <pr_pos name="transition_star_one" id="5564"
+       ip_theory="Vm">
+       <ip_library name="vm"/>
+       <ip_qualid name="transition_star_one"/>
+      </pr_pos>
+      <pr_pos name="transition_star_transitive" id="5571"
+       ip_theory="Vm">
+       <ip_library name="vm"/>
+       <ip_qualid name="transition_star_transitive"/>
+      </pr_pos>
+      <pr_pos name="seq_wp_lemma" id="5752"
+       ip_theory="Compiler_logic">
+       <ip_library name="logic"/>
+       <ip_qualid name="seq_wp_lemma"/>
+      </pr_pos>
+      <pr_pos name="fork_wp_lemma" id="5830"
+       ip_theory="Compiler_logic">
+       <ip_library name="logic"/>
+       <ip_qualid name="fork_wp_lemma"/>
+      </pr_pos>
+      <pr_pos name="towp_wp_lemma" id="5890"
+       ip_theory="Compiler_logic">
+       <ip_library name="logic"/>
+       <ip_qualid name="towp_wp_lemma"/>
+      </pr_pos>
+      <pr_pos name="trivial_pre_lemma" id="5963"
+       ip_theory="Compiler_logic">
+       <ip_library name="logic"/>
+       <ip_qualid name="trivial_pre_lemma"/>
+      </pr_pos>
+      <meta name="remove_logic">
+       <meta_arg_ls id="10"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="15"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="689"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="690"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="693"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="1859"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="1860"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="1861"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="1974"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="2099"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="2102"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="3648"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="3657"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="3671"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="4902"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="4903"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="4940"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="5329"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="5580"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="5984"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="5998"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="6034"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="6242"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="6341"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="6354"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="6401"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="6433"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="6496"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="6543"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="6556"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="6569"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="6582"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="6797"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="8537"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="8596"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="10482"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="10526"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="10605"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="10644"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="10716"/>
+      </meta>
+      <meta name="remove_logic">
+       <meta_arg_ls id="10789"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1862"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1869"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1872"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1875"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1878"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1881"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1886"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1893"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1900"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1918"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1923"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1926"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1938"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1941"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1948"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1953"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1958"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1959"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1979"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="1984"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="2105"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="2110"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="2115"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="2120"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="2123"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="2126"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="2131"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="2136"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="2139"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="2142"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="2145"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="2148"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="2151"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="2158"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="4292"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="5151"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="5564"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="5571"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="5752"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="5830"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="5890"/>
+      </meta>
+      <meta name="remove_prop">
+       <meta_arg_pr id="5963"/>
+      </meta>
+      <meta name="remove_type">
+       <meta_arg_ts id="2"/>
+      </meta>
+      <meta name="remove_type">
+       <meta_arg_ts id="8"/>
+      </meta>
+      <meta name="remove_type">
+       <meta_arg_ts id="54"/>
+      </meta>
+      <meta name="remove_type">
+       <meta_arg_ts id="4946"/>
+      </meta>
+      <meta name="remove_type">
+       <meta_arg_ts id="5227"/>
+      </meta>
+      <meta name="remove_type">
+       <meta_arg_ts id="5228"/>
+      </meta>
+      <meta name="remove_type">
+       <meta_arg_ts id="5244"/>
+      </meta>
+      <meta name="remove_type">
+       <meta_arg_ts id="5635"/>
+      </meta>
+      <meta name="remove_type">
+       <meta_arg_ts id="5636"/>
+      </meta>
+      <meta name="remove_type">
+       <meta_arg_ts id="6208"/>
+      </meta>
+      <meta name="remove_type">
+       <meta_arg_ts id="6495"/>
+      </meta>
+      <goal name="WP_parameter compile_com.49.1.1.6" expl="6.">
+      <transf name="eliminate_builtin">
+       <goal name="WP_parameter compile_com.49.1.1.6.1" expl="1.">
+       <proof prover="2" timelimit="30"><result status="valid" time="0.52"/></proof>
+       </goal>
+      </transf>
+      </goal>
+     </metas>
      </goal>
     </transf>
     </goal>
@@ -1059,28 +1816,28 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_com.50" expl="50. precondition">
-  <proof prover="3"><result status="valid" time="0.12"/></proof>
+  <proof prover="4"><result status="valid" time="0.12"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.51" expl="51. precondition">
-  <proof prover="1" edited="compiler_Compile_com_WP_parameter_compile_com_1.v"><result status="valid" time="10.90"/></proof>
+  <proof prover="1" edited="compiler_Compile_com_WP_parameter_compile_com_1.v"><result status="valid" time="11.94"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.52" expl="52. precondition">
-  <proof prover="3"><result status="valid" time="0.14"/></proof>
+  <proof prover="4"><result status="valid" time="0.14"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.53" expl="53. precondition">
-  <proof prover="3"><result status="valid" time="0.13"/></proof>
+  <proof prover="4"><result status="valid" time="0.13"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.54" expl="54. precondition">
-  <proof prover="3"><result status="valid" time="0.15"/></proof>
+  <proof prover="4"><result status="valid" time="0.15"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.55" expl="55. precondition">
-  <proof prover="3"><result status="valid" time="0.15"/></proof>
+  <proof prover="4"><result status="valid" time="0.15"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.56" expl="56. precondition">
-  <proof prover="3"><result status="valid" time="0.11"/></proof>
+  <proof prover="4"><result status="valid" time="0.11"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.57" expl="57. precondition">
-  <proof prover="3"><result status="valid" time="0.11"/></proof>
+  <proof prover="4"><result status="valid" time="0.11"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.58" expl="58. precondition">
   <transf name="simplify_trivial_quantification_in_goal">
@@ -1089,25 +1846,25 @@
     <goal name="WP_parameter compile_com.58.1.1" expl="1.">
     <transf name="split_goal_wp">
      <goal name="WP_parameter compile_com.58.1.1.1" expl="1.">
-     <proof prover="3"><result status="valid" time="0.45"/></proof>
+     <proof prover="4"><result status="valid" time="0.69"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.58.1.1.2" expl="2.">
-     <proof prover="3"><result status="valid" time="0.46"/></proof>
+     <proof prover="4"><result status="valid" time="0.70"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.58.1.1.3" expl="3.">
-     <proof prover="3"><result status="valid" time="0.48"/></proof>
+     <proof prover="4"><result status="valid" time="0.48"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.58.1.1.4" expl="4.">
-     <proof prover="3"><result status="valid" time="0.58"/></proof>
+     <proof prover="4"><result status="valid" time="0.58"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.58.1.1.5" expl="5.">
-     <proof prover="3"><result status="valid" time="0.46"/></proof>
+     <proof prover="4"><result status="valid" time="0.46"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.58.1.1.6" expl="6.">
-     <proof prover="3"><result status="valid" time="0.47"/></proof>
+     <proof prover="4"><result status="valid" time="0.46"/></proof>
      </goal>
      <goal name="WP_parameter compile_com.58.1.1.7" expl="7.">
-     <proof prover="3"><result status="valid" time="0.47"/></proof>
+     <proof prover="4"><result status="valid" time="0.47"/></proof>
      </goal>
     </transf>
     </goal>
@@ -1116,1288 +1873,48 @@
   </transf>
   </goal>
   <goal name="WP_parameter compile_com.59" expl="59. postcondition">
-  <proof prover="3"><result status="valid" time="0.14"/></proof>
+  <proof prover="4"><result status="valid" time="0.14"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.60" expl="60. postcondition">
-  <proof prover="3"><result status="valid" time="0.14"/></proof>
+  <proof prover="4"><result status="valid" time="0.14"/></proof>
   </goal>
   <goal name="WP_parameter compile_com.61" expl="61. postcondition">
-  <proof prover="3"><result status="valid" time="0.14"/></proof>
+  <proof prover="4"><result status="valid" time="0.14"/></proof>
   </goal>
  </transf>
  </goal>
- <goal name="WP_parameter compile_com_natural" expl="VC for compile_com_natural" expanded="true">
- <transf name="split_goal_wp" expanded="true">
+ <goal name="WP_parameter compile_com_natural" expl="VC for compile_com_natural">
+ <transf name="split_goal_wp">
   <goal name="WP_parameter compile_com_natural.1" expl="1. assertion">
-  <proof prover="3"><result status="valid" time="0.08"/></proof>
-  </goal>
-  <goal name="WP_parameter compile_com_natural.2" expl="2. postcondition" expanded="true">
-  <proof prover="1" edited="compiler_Compile_com_WP_parameter_compile_com_natural_1.v"><result status="valid" time="1.94"/></proof>
-  <metas
-   expanded="true">
-   <ts_pos name="real" arity="0" id="2"
-    ip_theory="BuiltIn">
-    <ip_library name="why3"/>
-    <ip_library name="BuiltIn"/>
-    <ip_qualid name="real"/>
-   </ts_pos>
-   <ts_pos name="pred" arity="1" id="8"
-    ip_theory="HighOrd">
-    <ip_library name="why3"/>
-    <ip_library name="HighOrd"/>
-    <ip_qualid name="pred"/>
-   </ts_pos>
-   <ts_pos name="unit" arity="0" id="21"
-    ip_theory="Unit">
-    <ip_library name="why3"/>
-    <ip_library name="Unit"/>
-    <ip_qualid name="unit"/>
-   </ts_pos>
-   <ts_pos name="&apos;mark" arity="0" id="54"
-    ip_theory="Mark">
-    <ip_library name="why3"/>
-    <ip_library name="Mark"/>
-    <ip_qualid name="&apos;mark"/>
-   </ts_pos>
-   <ts_pos name="state" arity="0" id="5044"
-    ip_theory="State">
-    <ip_library name="state"/>
-    <ip_qualid name="state"/>
-   </ts_pos>
-   <ts_pos name="pos" arity="0" id="5444"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="pos"/>
-   </ts_pos>
-   <ts_pos name="stack" arity="0" id="5445"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="stack"/>
-   </ts_pos>
-   <ts_pos name="code" arity="0" id="5461"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="code"/>
-   </ts_pos>
-   <ts_pos name="trans" arity="0" id="5760"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="trans"/>
-   </ts_pos>
-   <ts_pos name="pred" arity="0" id="5900"
-    ip_theory="Compiler_logic">
-    <ip_library name="logic"/>
-    <ip_qualid name="pred"/>
-   </ts_pos>
-   <ts_pos name="post" arity="0" id="5901"
-    ip_theory="Compiler_logic">
-    <ip_library name="logic"/>
-    <ip_qualid name="post"/>
-   </ts_pos>
-   <ts_pos name="wp" arity="1" id="5908"
-    ip_theory="Compiler_logic">
-    <ip_library name="logic"/>
-    <ip_qualid name="wp"/>
-   </ts_pos>
-   <ts_pos name="binop" arity="0" id="6463"
-    ip_theory="VM_arith_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="binop"/>
-   </ts_pos>
-   <ts_pos name="cond" arity="0" id="6800"
-    ip_theory="VM_bool_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="cond"/>
-   </ts_pos>
-   <ls_pos name="infix =" id="10"
-    ip_theory="BuiltIn">
-    <ip_library name="why3"/>
-    <ip_library name="BuiltIn"/>
-    <ip_qualid name="infix ="/>
-   </ls_pos>
-   <ls_pos name="infix @" id="15"
-    ip_theory="HighOrd">
-    <ip_library name="why3"/>
-    <ip_library name="HighOrd"/>
-    <ip_qualid name="infix @"/>
-   </ls_pos>
-   <ls_pos name="zero" id="786"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="zero"/>
-   </ls_pos>
-   <ls_pos name="one" id="787"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="one"/>
-   </ls_pos>
-   <ls_pos name="infix &lt;" id="788"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="infix &lt;"/>
-   </ls_pos>
-   <ls_pos name="infix &gt;" id="791"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="infix &gt;"/>
-   </ls_pos>
-   <ls_pos name="infix +" id="1957"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="infix +"/>
-   </ls_pos>
-   <ls_pos name="prefix -" id="1958"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="prefix -"/>
-   </ls_pos>
-   <ls_pos name="infix *" id="1959"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="infix *"/>
-   </ls_pos>
-   <ls_pos name="infix &gt;=" id="2027"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="infix &gt;="/>
-   </ls_pos>
-   <ls_pos name="abs" id="2072"
-    ip_theory="Abs">
-    <ip_library name="int"/>
-    <ip_qualid name="abs"/>
-   </ls_pos>
-   <ls_pos name="div" id="2197"
-    ip_theory="EuclideanDivision">
-    <ip_library name="int"/>
-    <ip_qualid name="div"/>
-   </ls_pos>
-   <ls_pos name="mod" id="2200"
-    ip_theory="EuclideanDivision">
-    <ip_library name="int"/>
-    <ip_qualid name="mod"/>
-   </ls_pos>
-   <ls_pos name="mem" id="2857"
-    ip_theory="Mem">
-    <ip_library name="list"/>
-    <ip_qualid name="mem"/>
-   </ls_pos>
-   <ls_pos name="orb" id="3746"
-    ip_theory="Bool">
-    <ip_library name="bool"/>
-    <ip_qualid name="orb"/>
-   </ls_pos>
-   <ls_pos name="xorb" id="3755"
-    ip_theory="Bool">
-    <ip_library name="bool"/>
-    <ip_qualid name="xorb"/>
-   </ls_pos>
-   <ls_pos name="implb" id="3769"
-    ip_theory="Bool">
-    <ip_library name="bool"/>
-    <ip_qualid name="implb"/>
-   </ls_pos>
-   <ls_pos name="get" id="5000"
-    ip_theory="State">
-    <ip_library name="state"/>
-    <ip_qualid name="get"/>
-   </ls_pos>
-   <ls_pos name="set" id="5001"
-    ip_theory="State">
-    <ip_library name="state"/>
-    <ip_qualid name="set"/>
-   </ls_pos>
-   <ls_pos name="const" id="5038"
-    ip_theory="State">
-    <ip_library name="state"/>
-    <ip_qualid name="const"/>
-   </ls_pos>
-   <ls_pos name="push" id="5494"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="push"/>
-   </ls_pos>
-   <ls_pos name="iconst" id="5503"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="iconst"/>
-   </ls_pos>
-   <ls_pos name="ivar" id="5508"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="ivar"/>
-   </ls_pos>
-   <ls_pos name="isetvar" id="5513"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="isetvar"/>
-   </ls_pos>
-   <ls_pos name="iadd" id="5518"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="iadd"/>
-   </ls_pos>
-   <ls_pos name="isub" id="5519"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="isub"/>
-   </ls_pos>
-   <ls_pos name="imul" id="5520"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="imul"/>
-   </ls_pos>
-   <ls_pos name="ibeq" id="5521"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="ibeq"/>
-   </ls_pos>
-   <ls_pos name="ible" id="5526"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="ible"/>
-   </ls_pos>
-   <ls_pos name="ibne" id="5531"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="ibne"/>
-   </ls_pos>
-   <ls_pos name="ibgt" id="5536"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="ibgt"/>
-   </ls_pos>
-   <ls_pos name="ibranch" id="5541"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="ibranch"/>
-   </ls_pos>
-   <ls_pos name="ihalt" id="5546"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="ihalt"/>
-   </ls_pos>
-   <ls_pos name="transition" id="5547"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="transition"/>
-   </ls_pos>
-   <ls_pos name="transition_star_proof" id="5769"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="transition_star_proof"/>
-   </ls_pos>
-   <ls_pos name="transition_star" id="5794"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="transition_star"/>
-   </ls_pos>
-   <ls_pos name="vm_terminates" id="5845"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="vm_terminates"/>
-   </ls_pos>
-   <ls_pos name="fst" id="5865"
-    ip_theory="Compiler_logic">
-    <ip_library name="logic"/>
-    <ip_qualid name="fst"/>
-   </ls_pos>
-   <ls_pos name="snd" id="5884"
-    ip_theory="Compiler_logic">
-    <ip_library name="logic"/>
-    <ip_qualid name="snd"/>
-   </ls_pos>
-   <ls_pos name="wp_correctness" id="5960"
-    ip_theory="Compiler_logic">
-    <ip_library name="logic"/>
-    <ip_qualid name="wp_correctness"/>
-   </ls_pos>
-   <ls_pos name="seq_wp" id="5986"
-    ip_theory="Compiler_logic">
-    <ip_library name="logic"/>
-    <ip_qualid name="seq_wp"/>
-   </ls_pos>
-   <ls_pos name="fork_wp" id="6064"
-    ip_theory="Compiler_logic">
-    <ip_library name="logic"/>
-    <ip_qualid name="fork_wp"/>
-   </ls_pos>
-   <ls_pos name="towp_wp" id="6125"
-    ip_theory="Compiler_logic">
-    <ip_library name="logic"/>
-    <ip_qualid name="towp_wp"/>
-   </ls_pos>
-   <ls_pos name="trivial_pre" id="6199"
-    ip_theory="Compiler_logic">
-    <ip_library name="logic"/>
-    <ip_qualid name="trivial_pre"/>
-   </ls_pos>
-   <ls_pos name="acc" id="6249"
-    ip_theory="Compiler_logic">
-    <ip_library name="logic"/>
-    <ip_qualid name="acc"/>
-   </ls_pos>
-   <ls_pos name="loop_preservation" id="6263"
-    ip_theory="Compiler_logic">
-    <ip_library name="logic"/>
-    <ip_qualid name="loop_preservation"/>
-   </ls_pos>
-   <ls_pos name="forget_old" id="6299"
-    ip_theory="Compiler_logic">
-    <ip_library name="logic"/>
-    <ip_qualid name="forget_old"/>
-   </ls_pos>
-   <ls_pos name="iconst_post" id="6339"
-    ip_theory="VM_arith_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="iconst_post"/>
-   </ls_pos>
-   <ls_pos name="ivar_post" id="6401"
-    ip_theory="VM_arith_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="ivar_post"/>
-   </ls_pos>
-   <ls_pos name="ibinop_pre" id="6464"
-    ip_theory="VM_arith_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="ibinop_pre"/>
-   </ls_pos>
-   <ls_pos name="ibinop_post" id="6519"
-    ip_theory="VM_arith_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="ibinop_post"/>
-   </ls_pos>
-   <ls_pos name="plus" id="6633"
-    ip_theory="VM_arith_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="plus"/>
-   </ls_pos>
-   <ls_pos name="sub" id="6646"
-    ip_theory="VM_arith_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="sub"/>
-   </ls_pos>
-   <ls_pos name="mul" id="6659"
-    ip_theory="VM_arith_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="mul"/>
-   </ls_pos>
-   <ls_pos name="inil_post" id="6706"
-    ip_theory="VM_bool_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="inil_post"/>
-   </ls_pos>
-   <ls_pos name="ibranch_post" id="6738"
-    ip_theory="VM_bool_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="ibranch_post"/>
-   </ls_pos>
-   <ls_pos name="icjump_post" id="6801"
-    ip_theory="VM_bool_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="icjump_post"/>
-   </ls_pos>
-   <ls_pos name="beq" id="6879"
-    ip_theory="VM_bool_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="beq"/>
-   </ls_pos>
-   <ls_pos name="bne" id="6899"
-    ip_theory="VM_bool_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="bne"/>
-   </ls_pos>
-   <ls_pos name="ble" id="6919"
-    ip_theory="VM_bool_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="ble"/>
-   </ls_pos>
-   <ls_pos name="bgt" id="6939"
-    ip_theory="VM_bool_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="bgt"/>
-   </ls_pos>
-   <ls_pos name="isetvar_pre" id="7064"
-    ip_theory="VM_bool_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="isetvar_pre"/>
-   </ls_pos>
-   <ls_pos name="isetvar_post" id="7112"
-    ip_theory="VM_bool_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="isetvar_post"/>
-   </ls_pos>
-   <ls_pos name="aexpr_post" id="7205"
-    ip_theory="Compile_aexpr">
-    <ip_qualid name="aexpr_post"/>
-   </ls_pos>
-   <ls_pos name="bexpr_post" id="8980"
-    ip_theory="Compile_bexpr">
-    <ip_qualid name="bexpr_post"/>
-   </ls_pos>
-   <ls_pos name="exn_bool" id="9039"
-    ip_theory="Compile_bexpr">
-    <ip_qualid name="exn_bool"/>
-   </ls_pos>
-   <ls_pos name="exn_bool_old" id="11141"
-    ip_theory="Compile_com">
-    <ip_qualid name="exn_bool_old"/>
-   </ls_pos>
-   <ls_pos name="loop_invariant" id="11180"
-    ip_theory="Compile_com">
-    <ip_qualid name="loop_invariant"/>
-   </ls_pos>
-   <ls_pos name="loop_post" id="11252"
-    ip_theory="Compile_com">
-    <ip_qualid name="loop_post"/>
-   </ls_pos>
-   <ls_pos name="loop_variant" id="11325"
-    ip_theory="Compile_com">
-    <ip_qualid name="loop_variant"/>
-   </ls_pos>
-   <pr_pos name="Assoc" id="1960"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="CommutativeGroup"/>
-    <ip_qualid name="Assoc"/>
-   </pr_pos>
-   <pr_pos name="Unit_def_l" id="1967"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="CommutativeGroup"/>
-    <ip_qualid name="Unit_def_l"/>
-   </pr_pos>
-   <pr_pos name="Unit_def_r" id="1970"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="CommutativeGroup"/>
-    <ip_qualid name="Unit_def_r"/>
-   </pr_pos>
-   <pr_pos name="Inv_def_l" id="1973"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="CommutativeGroup"/>
-    <ip_qualid name="Inv_def_l"/>
-   </pr_pos>
-   <pr_pos name="Inv_def_r" id="1976"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="CommutativeGroup"/>
-    <ip_qualid name="Inv_def_r"/>
-   </pr_pos>
-   <pr_pos name="Comm" id="1979"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="CommutativeGroup"/>
-    <ip_qualid name="Comm"/>
-    <ip_qualid name="Comm"/>
-   </pr_pos>
-   <pr_pos name="Assoc" id="1984"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="Assoc"/>
-    <ip_qualid name="Assoc"/>
-   </pr_pos>
-   <pr_pos name="Mul_distr_l" id="1991"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="Mul_distr_l"/>
-   </pr_pos>
-   <pr_pos name="Mul_distr_r" id="1998"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="Mul_distr_r"/>
-   </pr_pos>
-   <pr_pos name="Comm" id="2016"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="Comm"/>
-    <ip_qualid name="Comm"/>
-   </pr_pos>
-   <pr_pos name="Unitary" id="2021"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="Unitary"/>
-   </pr_pos>
-   <pr_pos name="NonTrivialRing" id="2024"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="NonTrivialRing"/>
-   </pr_pos>
-   <pr_pos name="Refl" id="2036"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="Refl"/>
-   </pr_pos>
-   <pr_pos name="Trans" id="2039"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="Trans"/>
-   </pr_pos>
-   <pr_pos name="Antisymm" id="2046"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="Antisymm"/>
-   </pr_pos>
-   <pr_pos name="Total" id="2051"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="Total"/>
-   </pr_pos>
-   <pr_pos name="ZeroLessOne" id="2056"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="ZeroLessOne"/>
-   </pr_pos>
-   <pr_pos name="CompatOrderAdd" id="2057"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="CompatOrderAdd"/>
-   </pr_pos>
-   <pr_pos name="CompatOrderMult" id="2064"
-    ip_theory="Int">
-    <ip_library name="int"/>
-    <ip_qualid name="CompatOrderMult"/>
-   </pr_pos>
-   <pr_pos name="Abs_le" id="2077"
-    ip_theory="Abs">
-    <ip_library name="int"/>
-    <ip_qualid name="Abs_le"/>
-   </pr_pos>
-   <pr_pos name="Abs_pos" id="2082"
-    ip_theory="Abs">
-    <ip_library name="int"/>
-    <ip_qualid name="Abs_pos"/>
-   </pr_pos>
-   <pr_pos name="Div_mod" id="2203"
-    ip_theory="EuclideanDivision">
-    <ip_library name="int"/>
-    <ip_qualid name="Div_mod"/>
-   </pr_pos>
-   <pr_pos name="Div_bound" id="2208"
-    ip_theory="EuclideanDivision">
-    <ip_library name="int"/>
-    <ip_qualid name="Div_bound"/>
-   </pr_pos>
-   <pr_pos name="Mod_bound" id="2213"
-    ip_theory="EuclideanDivision">
-    <ip_library name="int"/>
-    <ip_qualid name="Mod_bound"/>
-   </pr_pos>
-   <pr_pos name="Mod_1" id="2218"
-    ip_theory="EuclideanDivision">
-    <ip_library name="int"/>
-    <ip_qualid name="Mod_1"/>
-   </pr_pos>
-   <pr_pos name="Div_1" id="2221"
-    ip_theory="EuclideanDivision">
-    <ip_library name="int"/>
-    <ip_qualid name="Div_1"/>
-   </pr_pos>
-   <pr_pos name="Div_inf" id="2224"
-    ip_theory="EuclideanDivision">
-    <ip_library name="int"/>
-    <ip_qualid name="Div_inf"/>
-   </pr_pos>
-   <pr_pos name="Div_inf_neg" id="2229"
-    ip_theory="EuclideanDivision">
-    <ip_library name="int"/>
-    <ip_qualid name="Div_inf_neg"/>
-   </pr_pos>
-   <pr_pos name="Mod_0" id="2234"
-    ip_theory="EuclideanDivision">
-    <ip_library name="int"/>
-    <ip_qualid name="Mod_0"/>
-   </pr_pos>
-   <pr_pos name="Div_1_left" id="2237"
-    ip_theory="EuclideanDivision">
-    <ip_library name="int"/>
-    <ip_qualid name="Div_1_left"/>
-   </pr_pos>
-   <pr_pos name="Div_minus1_left" id="2240"
-    ip_theory="EuclideanDivision">
-    <ip_library name="int"/>
-    <ip_qualid name="Div_minus1_left"/>
-   </pr_pos>
-   <pr_pos name="Mod_1_left" id="2243"
-    ip_theory="EuclideanDivision">
-    <ip_library name="int"/>
-    <ip_qualid name="Mod_1_left"/>
-   </pr_pos>
-   <pr_pos name="Mod_minus1_left" id="2246"
-    ip_theory="EuclideanDivision">
-    <ip_library name="int"/>
-    <ip_qualid name="Mod_minus1_left"/>
-   </pr_pos>
-   <pr_pos name="Div_mult" id="2249"
-    ip_theory="EuclideanDivision">
-    <ip_library name="int"/>
-    <ip_qualid name="Div_mult"/>
-   </pr_pos>
-   <pr_pos name="Mod_mult" id="2256"
-    ip_theory="EuclideanDivision">
-    <ip_library name="int"/>
-    <ip_qualid name="Mod_mult"/>
-   </pr_pos>
-   <pr_pos name="Length_nonnegative" id="2850"
-    ip_theory="Length">
-    <ip_library name="list"/>
-    <ip_qualid name="Length_nonnegative"/>
-   </pr_pos>
-   <pr_pos name="Length_nil" id="2853"
-    ip_theory="Length">
-    <ip_library name="list"/>
-    <ip_qualid name="Length_nil"/>
-   </pr_pos>
-   <pr_pos name="Append_assoc" id="4390"
-    ip_theory="Append">
-    <ip_library name="list"/>
-    <ip_qualid name="Append_assoc"/>
-   </pr_pos>
-   <pr_pos name="Append_l_nil" id="4397"
-    ip_theory="Append">
-    <ip_library name="list"/>
-    <ip_qualid name="Append_l_nil"/>
-   </pr_pos>
-   <pr_pos name="Append_length" id="4400"
-    ip_theory="Append">
-    <ip_library name="list"/>
-    <ip_qualid name="Append_length"/>
-   </pr_pos>
-   <pr_pos name="mem_append" id="4405"
-    ip_theory="Append">
-    <ip_library name="list"/>
-    <ip_qualid name="mem_append"/>
-   </pr_pos>
-   <pr_pos name="mem_decomp" id="4412"
-    ip_theory="Append">
-    <ip_library name="list"/>
-    <ip_qualid name="mem_decomp"/>
-   </pr_pos>
-   <pr_pos name="Select_eq" id="5026"
-    ip_theory="State">
-    <ip_library name="state"/>
-    <ip_qualid name="Select_eq"/>
-   </pr_pos>
-   <pr_pos name="Select_neq" id="5032"
-    ip_theory="State">
-    <ip_library name="state"/>
-    <ip_qualid name="Select_neq"/>
-   </pr_pos>
-   <pr_pos name="Const" id="5039"
-    ip_theory="State">
-    <ip_library name="state"/>
-    <ip_qualid name="Const"/>
-   </pr_pos>
-   <pr_pos name="inversion_beval_t" id="5157"
-    ip_theory="Imp">
-    <ip_library name="imp"/>
-    <ip_qualid name="inversion_beval_t"/>
-   </pr_pos>
-   <pr_pos name="inversion_beval_f" id="5164"
-    ip_theory="Imp">
-    <ip_library name="imp"/>
-    <ip_qualid name="inversion_beval_f"/>
-   </pr_pos>
-   <pr_pos name="codeseq_at_app_right" id="5476"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="codeseq_at_app_right"/>
-   </pr_pos>
-   <pr_pos name="codeseq_at_app_left" id="5485"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="codeseq_at_app_left"/>
-   </pr_pos>
-   <pr_pos name="transition_star_def" id="5795"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="transition_star_def"/>
-   </pr_pos>
-   <pr_pos name="transZero" id="5804"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="transZero"/>
-   </pr_pos>
-   <pr_pos name="transOne" id="5809"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="transOne"/>
-   </pr_pos>
-   <pr_pos name="transition_star_one" id="5818"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="transition_star_one"/>
-   </pr_pos>
-   <pr_pos name="trans_star" id="5825"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="trans_star"/>
-   </pr_pos>
-   <pr_pos name="transition_star_transitive" id="5836"
-    ip_theory="Vm">
-    <ip_library name="vm"/>
-    <ip_qualid name="transition_star_transitive"/>
-   </pr_pos>
-   <pr_pos name="seq_wp_lemma" id="6017"
-    ip_theory="Compiler_logic">
-    <ip_library name="logic"/>
-    <ip_qualid name="seq_wp_lemma"/>
-   </pr_pos>
-   <pr_pos name="fork_wp_lemma" id="6095"
-    ip_theory="Compiler_logic">
-    <ip_library name="logic"/>
-    <ip_qualid name="fork_wp_lemma"/>
-   </pr_pos>
-   <pr_pos name="towp_wp_lemma" id="6155"
-    ip_theory="Compiler_logic">
-    <ip_library name="logic"/>
-    <ip_qualid name="towp_wp_lemma"/>
-   </pr_pos>
-   <pr_pos name="trivial_pre_lemma" id="6228"
-    ip_theory="Compiler_logic">
-    <ip_library name="logic"/>
-    <ip_qualid name="trivial_pre_lemma"/>
-   </pr_pos>
-   <pr_pos name="ibinop_pre_lemma" id="6497"
-    ip_theory="VM_arith_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="ibinop_pre_lemma"/>
-   </pr_pos>
-   <pr_pos name="ibinop_post_lemma" id="6561"
-    ip_theory="VM_arith_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="ibinop_post_lemma"/>
-   </pr_pos>
-   <pr_pos name="icjump_post_lemma" id="6848"
-    ip_theory="VM_bool_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="icjump_post_lemma"/>
-   </pr_pos>
-   <pr_pos name="beq_lemma" id="6892"
-    ip_theory="VM_bool_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="beq_lemma"/>
-   </pr_pos>
-   <pr_pos name="bne_lemma" id="6912"
-    ip_theory="VM_bool_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="bne_lemma"/>
-   </pr_pos>
-   <pr_pos name="ble_lemma" id="6932"
-    ip_theory="VM_bool_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="ble_lemma"/>
-   </pr_pos>
-   <pr_pos name="bgt_lemma" id="6952"
-    ip_theory="VM_bool_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="bgt_lemma"/>
-   </pr_pos>
-   <pr_pos name="isetvar_pre_lemma" id="7093"
-    ip_theory="VM_bool_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="isetvar_pre_lemma"/>
-   </pr_pos>
-   <pr_pos name="isetvar_post_lemma" id="7150"
-    ip_theory="VM_bool_instr_spec">
-    <ip_library name="specs"/>
-    <ip_qualid name="isetvar_post_lemma"/>
-   </pr_pos>
-   <pr_pos name="aexpr_post_lemma" id="7254"
-    ip_theory="Compile_aexpr">
-    <ip_qualid name="aexpr_post_lemma"/>
-   </pr_pos>
-   <pr_pos name="com_pre_lemma" id="10969"
-    ip_theory="Compile_com">
-    <ip_qualid name="com_pre_lemma"/>
-   </pr_pos>
-   <pr_pos name="com_post_lemma" id="11080"
-    ip_theory="Compile_com">
-    <ip_qualid name="com_post_lemma"/>
-   </pr_pos>
-   <meta name="remove_logic">
-    <meta_arg_ls id="10"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="15"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="786"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="787"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="788"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="791"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="1957"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="1958"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="1959"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="2027"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="2072"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="2197"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="2200"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="2857"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="3746"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="3755"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="3769"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5000"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5001"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5038"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5494"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5503"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5508"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5513"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5518"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5519"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5520"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5521"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5526"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5531"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5536"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5541"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5546"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5547"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5769"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5794"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5845"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5865"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5884"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5960"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="5986"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6064"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6125"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6199"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6249"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6263"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6299"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6339"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6401"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6464"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6519"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6633"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6646"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6659"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6706"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6738"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6801"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6879"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6899"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6919"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="6939"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="7064"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="7112"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="7205"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="8980"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="9039"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="11141"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="11180"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="11252"/>
-   </meta>
-   <meta name="remove_logic">
-    <meta_arg_ls id="11325"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="1960"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="1967"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="1970"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="1973"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="1976"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="1979"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="1984"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="1991"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="1998"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2016"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2021"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2024"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2036"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2039"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2046"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2051"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2056"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2057"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2064"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2077"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2082"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2203"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2208"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2213"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2218"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2221"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2224"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2229"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2234"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2237"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2240"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2243"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2246"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2249"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2256"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2850"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="2853"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="4390"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="4397"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="4400"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="4405"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="4412"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="5026"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="5032"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="5039"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="5157"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="5164"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="5476"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="5485"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="5795"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="5804"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="5809"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="5818"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="5825"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="5836"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="6017"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="6095"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="6155"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="6228"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="6497"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="6561"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="6848"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="6892"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="6912"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="6932"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="6952"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="7093"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="7150"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="7254"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="10969"/>
-   </meta>
-   <meta name="remove_prop">
-    <meta_arg_pr id="11080"/>
-   </meta>
-   <meta name="remove_type">
-    <meta_arg_ts id="2"/>
-   </meta>
-   <meta name="remove_type">
-    <meta_arg_ts id="8"/>
-   </meta>
-   <meta name="remove_type">
-    <meta_arg_ts id="21"/>
-   </meta>
-   <meta name="remove_type">
-    <meta_arg_ts id="54"/>
-   </meta>
-   <meta name="remove_type">
-    <meta_arg_ts id="5044"/>
-   </meta>
-   <meta name="remove_type">
-    <meta_arg_ts id="5444"/>
-   </meta>
-   <meta name="remove_type">
-    <meta_arg_ts id="5445"/>
-   </meta>
-   <meta name="remove_type">
-    <meta_arg_ts id="5461"/>
-   </meta>
-   <meta name="remove_type">
-    <meta_arg_ts id="5760"/>
-   </meta>
-   <meta name="remove_type">
-    <meta_arg_ts id="5900"/>
-   </meta>
-   <meta name="remove_type">
-    <meta_arg_ts id="5901"/>
-   </meta>
-   <meta name="remove_type">
-    <meta_arg_ts id="5908"/>
-   </meta>
-   <meta name="remove_type">
-    <meta_arg_ts id="6463"/>
-   </meta>
-   <meta name="remove_type">
-    <meta_arg_ts id="6800"/>
-   </meta>
-   <goal name="WP_parameter compile_com_natural.2" expl="2. postcondition" expanded="true">
-   <transf name="eliminate_builtin" expanded="true">
-    <goal name="WP_parameter compile_com_natural.2.1" expl="1. postcondition" expanded="true">
-    <proof prover="1" edited="compiler_Compile_com_WP_parameter_compile_com_natural_2.v"><result status="valid" time="1.30"/></proof>
+  <transf name="split_goal_wp">
+   <goal name="WP_parameter compile_com_natural.1.1" expl="1. assertion">
+   <proof prover="4"><result status="valid" time="0.09"/></proof>
+   </goal>
+   <goal name="WP_parameter compile_com_natural.1.2" expl="2. assertion">
+   <transf name="simplify_trivial_quantification_in_goal">
+    <goal name="WP_parameter compile_com_natural.1.2.1" expl="1.">
+    <transf name="compute_specified">
+     <goal name="WP_parameter compile_com_natural.1.2.1.1" expl="1.">
+     <proof prover="2"><result status="valid" time="0.05"/></proof>
+     <proof prover="4"><result status="valid" time="0.08"/></proof>
+     </goal>
+    </transf>
     </goal>
    </transf>
    </goal>
-  </metas>
+  </transf>
+  </goal>
+  <goal name="WP_parameter compile_com_natural.2" expl="2. postcondition">
+  <proof prover="4"><result status="valid" time="0.10"/></proof>
   </goal>
  </transf>
  </goal>
- <goal name="WP_parameter compile_program" expl="VC for compile_program" expanded="true">
- <transf name="split_goal_wp" expanded="true">
-  <goal name="WP_parameter compile_program.1" expl="1. postcondition" expanded="true">
-  <transf name="inline_goal" expanded="true">
-   <goal name="WP_parameter compile_program.1.1" expl="1. postcondition" expanded="true">
-   <metas
-    expanded="true">
+ <goal name="WP_parameter compile_program" expl="VC for compile_program">
+ <transf name="split_goal_wp">
+  <goal name="WP_parameter compile_program.1" expl="1. postcondition">
+  <transf name="inline_goal">
+   <goal name="WP_parameter compile_program.1.1" expl="1. postcondition">
+   <metas>
     <ts_pos name="real" arity="0" id="2"
      ip_theory="BuiltIn">
      <ip_library name="why3"/>
@@ -2428,63 +1945,58 @@
      <ip_library name="Mark"/>
      <ip_qualid name="&apos;mark"/>
     </ts_pos>
-    <ts_pos name="tuple2" arity="2" id="1502"
+    <ts_pos name="tuple2" arity="2" id="1404"
      ip_theory="Tuple2">
      <ip_library name="why3"/>
      <ip_library name="Tuple2"/>
      <ip_qualid name="tuple2"/>
     </ts_pos>
-    <ts_pos name="state" arity="0" id="5044"
+    <ts_pos name="state" arity="0" id="4946"
      ip_theory="State">
      <ip_library name="state"/>
      <ip_qualid name="state"/>
     </ts_pos>
-    <ts_pos name="pos" arity="0" id="5444"
+    <ts_pos name="pos" arity="0" id="5227"
      ip_theory="Vm">
      <ip_library name="vm"/>
      <ip_qualid name="pos"/>
     </ts_pos>
-    <ts_pos name="stack" arity="0" id="5445"
+    <ts_pos name="stack" arity="0" id="5228"
      ip_theory="Vm">
      <ip_library name="vm"/>
      <ip_qualid name="stack"/>
     </ts_pos>
-    <ts_pos name="code" arity="0" id="5461"
+    <ts_pos name="code" arity="0" id="5244"
      ip_theory="Vm">
      <ip_library name="vm"/>
      <ip_qualid name="code"/>
     </ts_pos>
-    <ts_pos name="trans" arity="0" id="5760"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="trans"/>
-    </ts_pos>
-    <ts_pos name="pred" arity="0" id="5900"
+    <ts_pos name="pred" arity="0" id="5635"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="pred"/>
     </ts_pos>
-    <ts_pos name="post" arity="0" id="5901"
+    <ts_pos name="post" arity="0" id="5636"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="post"/>
     </ts_pos>
-    <ts_pos name="hl" arity="1" id="5902"
+    <ts_pos name="hl" arity="1" id="5637"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="hl"/>
     </ts_pos>
-    <ts_pos name="wp" arity="1" id="5908"
+    <ts_pos name="wp" arity="1" id="5643"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="wp"/>
     </ts_pos>
-    <ts_pos name="binop" arity="0" id="6463"
+    <ts_pos name="binop" arity="0" id="6208"
      ip_theory="VM_arith_instr_spec">
      <ip_library name="specs"/>
      <ip_qualid name="binop"/>
     </ts_pos>
-    <ts_pos name="cond" arity="0" id="6800"
+    <ts_pos name="cond" arity="0" id="6495"
      ip_theory="VM_bool_instr_spec">
      <ip_library name="specs"/>
      <ip_qualid name="cond"/>
@@ -2501,702 +2013,545 @@
      <ip_library name="HighOrd"/>
      <ip_qualid name="infix @"/>
     </ls_pos>
-    <ls_pos name="one" id="787"
+    <ls_pos name="one" id="689"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="one"/>
     </ls_pos>
-    <ls_pos name="infix &lt;" id="788"
+    <ls_pos name="infix &lt;" id="690"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="infix &lt;"/>
     </ls_pos>
-    <ls_pos name="infix &gt;" id="791"
+    <ls_pos name="infix &gt;" id="693"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="infix &gt;"/>
     </ls_pos>
-    <ls_pos name="infix +" id="1957"
+    <ls_pos name="infix +" id="1859"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="infix +"/>
     </ls_pos>
-    <ls_pos name="prefix -" id="1958"
+    <ls_pos name="prefix -" id="1860"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="prefix -"/>
     </ls_pos>
-    <ls_pos name="infix *" id="1959"
+    <ls_pos name="infix *" id="1861"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="infix *"/>
     </ls_pos>
-    <ls_pos name="infix &gt;=" id="2027"
+    <ls_pos name="infix &gt;=" id="1929"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="infix &gt;="/>
     </ls_pos>
-    <ls_pos name="abs" id="2072"
+    <ls_pos name="abs" id="1974"
      ip_theory="Abs">
      <ip_library name="int"/>
      <ip_qualid name="abs"/>
     </ls_pos>
-    <ls_pos name="div" id="2197"
+    <ls_pos name="div" id="2099"
      ip_theory="EuclideanDivision">
      <ip_library name="int"/>
      <ip_qualid name="div"/>
     </ls_pos>
-    <ls_pos name="mod" id="2200"
+    <ls_pos name="mod" id="2102"
      ip_theory="EuclideanDivision">
      <ip_library name="int"/>
      <ip_qualid name="mod"/>
     </ls_pos>
-    <ls_pos name="mem" id="2857"
+    <ls_pos name="mem" id="2759"
      ip_theory="Mem">
      <ip_library name="list"/>
      <ip_qualid name="mem"/>
     </ls_pos>
-    <ls_pos name="orb" id="3746"
+    <ls_pos name="orb" id="3648"
      ip_theory="Bool">
      <ip_library name="bool"/>
      <ip_qualid name="orb"/>
     </ls_pos>
-    <ls_pos name="xorb" id="3755"
+    <ls_pos name="xorb" id="3657"
      ip_theory="Bool">
      <ip_library name="bool"/>
      <ip_qualid name="xorb"/>
     </ls_pos>
-    <ls_pos name="implb" id="3769"
+    <ls_pos name="implb" id="3671"
      ip_theory="Bool">
      <ip_library name="bool"/>
      <ip_qualid name="implb"/>
     </ls_pos>
-    <ls_pos name="get" id="5000"
+    <ls_pos name="get" id="4902"
      ip_theory="State">
      <ip_library name="state"/>
      <ip_qualid name="get"/>
     </ls_pos>
-    <ls_pos name="set" id="5001"
+    <ls_pos name="set" id="4903"
      ip_theory="State">
      <ip_library name="state"/>
      <ip_qualid name="set"/>
     </ls_pos>
-    <ls_pos name="const" id="5038"
+    <ls_pos name="const" id="4940"
      ip_theory="State">
      <ip_library name="state"/>
      <ip_qualid name="const"/>
     </ls_pos>
-    <ls_pos name="push" id="5494"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="push"/>
-    </ls_pos>
-    <ls_pos name="iconst" id="5503"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="iconst"/>
-    </ls_pos>
-    <ls_pos name="ivar" id="5508"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="ivar"/>
-    </ls_pos>
-    <ls_pos name="isetvar" id="5513"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="isetvar"/>
-    </ls_pos>
-    <ls_pos name="iadd" id="5518"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="iadd"/>
-    </ls_pos>
-    <ls_pos name="isub" id="5519"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="isub"/>
-    </ls_pos>
-    <ls_pos name="imul" id="5520"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="imul"/>
-    </ls_pos>
-    <ls_pos name="ibeq" id="5521"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="ibeq"/>
-    </ls_pos>
-    <ls_pos name="ible" id="5526"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="ible"/>
-    </ls_pos>
-    <ls_pos name="ibne" id="5531"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="ibne"/>
-    </ls_pos>
-    <ls_pos name="ibgt" id="5536"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="ibgt"/>
-    </ls_pos>
-    <ls_pos name="ibranch" id="5541"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="ibranch"/>
-    </ls_pos>
-    <ls_pos name="transition" id="5547"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="transition"/>
-    </ls_pos>
-    <ls_pos name="transition_star_proof" id="5769"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="transition_star_proof"/>
-    </ls_pos>
-    <ls_pos name="transition_star" id="5794"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="transition_star"/>
-    </ls_pos>
-    <ls_pos name="vm_terminates" id="5845"
+    <ls_pos name="vm_terminates" id="5580"
      ip_theory="Vm">
      <ip_library name="vm"/>
      <ip_qualid name="vm_terminates"/>
     </ls_pos>
-    <ls_pos name="fst" id="5865"
+    <ls_pos name="fst" id="5600"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="fst"/>
     </ls_pos>
-    <ls_pos name="snd" id="5884"
+    <ls_pos name="snd" id="5619"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="snd"/>
     </ls_pos>
-    <ls_pos name="contextual_irrelevance" id="5913"
+    <ls_pos name="contextual_irrelevance" id="5648"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="contextual_irrelevance"/>
     </ls_pos>
-    <ls_pos name="hl_correctness" id="5938"
+    <ls_pos name="hl_correctness" id="5673"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="hl_correctness"/>
     </ls_pos>
-    <ls_pos name="wp_correctness" id="5960"
+    <ls_pos name="wp_correctness" id="5695"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="wp_correctness"/>
     </ls_pos>
-    <ls_pos name="seq_wp" id="5986"
+    <ls_pos name="seq_wp" id="5721"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="seq_wp"/>
     </ls_pos>
-    <ls_pos name="fork_wp" id="6064"
+    <ls_pos name="fork_wp" id="5799"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="fork_wp"/>
     </ls_pos>
-    <ls_pos name="towp_wp" id="6125"
+    <ls_pos name="towp_wp" id="5860"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="towp_wp"/>
     </ls_pos>
-    <ls_pos name="trivial_pre" id="6199"
+    <ls_pos name="trivial_pre" id="5934"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="trivial_pre"/>
     </ls_pos>
-    <ls_pos name="acc" id="6249"
+    <ls_pos name="acc" id="5984"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="acc"/>
     </ls_pos>
-    <ls_pos name="loop_preservation" id="6263"
+    <ls_pos name="loop_preservation" id="5998"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="loop_preservation"/>
     </ls_pos>
-    <ls_pos name="forget_old" id="6299"
+    <ls_pos name="forget_old" id="6034"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="forget_old"/>
     </ls_pos>
-    <ls_pos name="iconst_post" id="6339"
+    <ls_pos name="iconst_post" id="6084"
      ip_theory="VM_arith_instr_spec">
      <ip_library name="specs"/>
      <ip_qualid name="iconst_post"/>
     </ls_pos>
-    <ls_pos name="ivar_post" id="6401"
+    <ls_pos name="ivar_post" id="6146"
      ip_theory="VM_arith_instr_spec">
      <ip_library name="specs"/>
      <ip_qualid name="ivar_post"/>
     </ls_pos>
-    <ls_pos name="ibinop_pre" id="6464"
+    <ls_pos name="ibinop_pre" id="6209"
      ip_theory="VM_arith_instr_spec">
      <ip_library name="specs"/>
      <ip_qualid name="ibinop_pre"/>
     </ls_pos>
-    <ls_pos name="ibinop_post" id="6519"
+    <ls_pos name="ibinop_post" id="6242"
      ip_theory="VM_arith_instr_spec">
      <ip_library name="specs"/>
      <ip_qualid name="ibinop_post"/>
     </ls_pos>
-    <ls_pos name="plus" id="6633"
+    <ls_pos name="plus" id="6328"
      ip_theory="VM_arith_instr_spec">
      <ip_library name="specs"/>
      <ip_qualid name="plus"/>
     </ls_pos>
-    <ls_pos name="sub" id="6646"
+    <ls_pos name="sub" id="6341"
      ip_theory="VM_arith_instr_spec">
      <ip_library name="specs"/>
      <ip_qualid name="sub"/>
     </ls_pos>
-    <ls_pos name="mul" id="6659"
+    <ls_pos name="mul" id="6354"
      ip_theory="VM_arith_instr_spec">
      <ip_library name="specs"/>
      <ip_qualid name="mul"/>
     </ls_pos>
-    <ls_pos name="inil_post" id="6706"
+    <ls_pos name="inil_post" id="6401"
      ip_theory="VM_bool_instr_spec">
      <ip_library name="specs"/>
      <ip_qualid name="inil_post"/>
     </ls_pos>
-    <ls_pos name="ibranch_post" id="6738"
+    <ls_pos name="ibranch_post" id="6433"
      ip_theory="VM_bool_instr_spec">
      <ip_library name="specs"/>
      <ip_qualid name="ibranch_post"/>
     </ls_pos>
-    <ls_pos name="icjump_post" id="6801"
+    <ls_pos name="icjump_post" id="6496"
      ip_theory="VM_bool_instr_spec">
      <ip_library name="specs"/>
      <ip_qualid name="icjump_post"/>
     </ls_pos>
-    <ls_pos name="beq" id="6879"
+    <ls_pos name="beq" id="6543"
      ip_theory="VM_bool_instr_spec">
      <ip_library name="specs"/>
      <ip_qualid name="beq"/>
     </ls_pos>
-    <ls_pos name="bne" id="6899"
+    <ls_pos name="bne" id="6556"
      ip_theory="VM_bool_instr_spec">
      <ip_library name="specs"/>
      <ip_qualid name="bne"/>
     </ls_pos>
-    <ls_pos name="ble" id="6919"
+    <ls_pos name="ble" id="6569"
      ip_theory="VM_bool_instr_spec">
      <ip_library name="specs"/>
      <ip_qualid name="ble"/>
     </ls_pos>
-    <ls_pos name="bgt" id="6939"
+    <ls_pos name="bgt" id="6582"
      ip_theory="VM_bool_instr_spec">
      <ip_library name="specs"/>
      <ip_qualid name="bgt"/>
     </ls_pos>
-    <ls_pos name="isetvar_pre" id="7064"
+    <ls_pos name="isetvar_pre" id="6700"
      ip_theory="VM_bool_instr_spec">
      <ip_library name="specs"/>
      <ip_qualid name="isetvar_pre"/>
     </ls_pos>
-    <ls_pos name="isetvar_post" id="7112"
+    <ls_pos name="isetvar_post" id="6729"
      ip_theory="VM_bool_instr_spec">
      <ip_library name="specs"/>
      <ip_qualid name="isetvar_post"/>
     </ls_pos>
-    <ls_pos name="aexpr_post" id="7205"
+    <ls_pos name="aexpr_post" id="6797"
      ip_theory="Compile_aexpr">
      <ip_qualid name="aexpr_post"/>
     </ls_pos>
-    <ls_pos name="bexpr_post" id="8980"
+    <ls_pos name="bexpr_post" id="8537"
      ip_theory="Compile_bexpr">
      <ip_qualid name="bexpr_post"/>
     </ls_pos>
-    <ls_pos name="exn_bool" id="9039"
+    <ls_pos name="exn_bool" id="8596"
      ip_theory="Compile_bexpr">
      <ip_qualid name="exn_bool"/>
     </ls_pos>
-    <ls_pos name="com_pre" id="10925"
+    <ls_pos name="com_pre" id="10482"
      ip_theory="Compile_com">
      <ip_qualid name="com_pre"/>
     </ls_pos>
-    <ls_pos name="com_post" id="11001"
+    <ls_pos name="com_post" id="10526"
      ip_theory="Compile_com">
      <ip_qualid name="com_post"/>
     </ls_pos>
-    <ls_pos name="exn_bool_old" id="11141"
+    <ls_pos name="exn_bool_old" id="10605"
      ip_theory="Compile_com">
      <ip_qualid name="exn_bool_old"/>
     </ls_pos>
-    <ls_pos name="loop_invariant" id="11180"
+    <ls_pos name="loop_invariant" id="10644"
      ip_theory="Compile_com">
      <ip_qualid name="loop_invariant"/>
     </ls_pos>
-    <ls_pos name="loop_post" id="11252"
+    <ls_pos name="loop_post" id="10716"
      ip_theory="Compile_com">
      <ip_qualid name="loop_post"/>
     </ls_pos>
-    <ls_pos name="loop_variant" id="11325"
+    <ls_pos name="loop_variant" id="10789"
      ip_theory="Compile_com">
      <ip_qualid name="loop_variant"/>
     </ls_pos>
-    <pr_pos name="Assoc" id="1960"
+    <pr_pos name="Assoc" id="1862"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="CommutativeGroup"/>
      <ip_qualid name="Assoc"/>
     </pr_pos>
-    <pr_pos name="Unit_def_r" id="1970"
+    <pr_pos name="Unit_def_r" id="1872"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="CommutativeGroup"/>
      <ip_qualid name="Unit_def_r"/>
     </pr_pos>
-    <pr_pos name="Inv_def_l" id="1973"
+    <pr_pos name="Inv_def_l" id="1875"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="CommutativeGroup"/>
      <ip_qualid name="Inv_def_l"/>
     </pr_pos>
-    <pr_pos name="Inv_def_r" id="1976"
+    <pr_pos name="Inv_def_r" id="1878"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="CommutativeGroup"/>
      <ip_qualid name="Inv_def_r"/>
     </pr_pos>
-    <pr_pos name="Comm" id="1979"
+    <pr_pos name="Comm" id="1881"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="CommutativeGroup"/>
      <ip_qualid name="Comm"/>
      <ip_qualid name="Comm"/>
     </pr_pos>
-    <pr_pos name="Assoc" id="1984"
+    <pr_pos name="Assoc" id="1886"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="Assoc"/>
      <ip_qualid name="Assoc"/>
     </pr_pos>
-    <pr_pos name="Mul_distr_l" id="1991"
+    <pr_pos name="Mul_distr_l" id="1893"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="Mul_distr_l"/>
     </pr_pos>
-    <pr_pos name="Mul_distr_r" id="1998"
+    <pr_pos name="Mul_distr_r" id="1900"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="Mul_distr_r"/>
     </pr_pos>
-    <pr_pos name="Comm" id="2016"
+    <pr_pos name="Comm" id="1918"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="Comm"/>
      <ip_qualid name="Comm"/>
     </pr_pos>
-    <pr_pos name="Unitary" id="2021"
+    <pr_pos name="Unitary" id="1923"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="Unitary"/>
     </pr_pos>
-    <pr_pos name="NonTrivialRing" id="2024"
+    <pr_pos name="NonTrivialRing" id="1926"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="NonTrivialRing"/>
     </pr_pos>
-    <pr_pos name="Refl" id="2036"
+    <pr_pos name="Refl" id="1938"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="Refl"/>
     </pr_pos>
-    <pr_pos name="Trans" id="2039"
+    <pr_pos name="Trans" id="1941"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="Trans"/>
     </pr_pos>
-    <pr_pos name="Antisymm" id="2046"
+    <pr_pos name="Antisymm" id="1948"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="Antisymm"/>
     </pr_pos>
-    <pr_pos name="Total" id="2051"
+    <pr_pos name="Total" id="1953"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="Total"/>
     </pr_pos>
-    <pr_pos name="ZeroLessOne" id="2056"
+    <pr_pos name="ZeroLessOne" id="1958"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="ZeroLessOne"/>
     </pr_pos>
-    <pr_pos name="CompatOrderAdd" id="2057"
+    <pr_pos name="CompatOrderAdd" id="1959"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="CompatOrderAdd"/>
     </pr_pos>
-    <pr_pos name="CompatOrderMult" id="2064"
+    <pr_pos name="CompatOrderMult" id="1966"
      ip_theory="Int">
      <ip_library name="int"/>
      <ip_qualid name="CompatOrderMult"/>
     </pr_pos>
-    <pr_pos name="Abs_le" id="2077"
+    <pr_pos name="Abs_le" id="1979"
      ip_theory="Abs">
      <ip_library name="int"/>
      <ip_qualid name="Abs_le"/>
     </pr_pos>
-    <pr_pos name="Abs_pos" id="2082"
+    <pr_pos name="Abs_pos" id="1984"
      ip_theory="Abs">
      <ip_library name="int"/>
      <ip_qualid name="Abs_pos"/>
     </pr_pos>
-    <pr_pos name="Div_mod" id="2203"
+    <pr_pos name="Div_mod" id="2105"
      ip_theory="EuclideanDivision">
      <ip_library name="int"/>
      <ip_qualid name="Div_mod"/>
     </pr_pos>
-    <pr_pos name="Div_bound" id="2208"
+    <pr_pos name="Div_bound" id="2110"
      ip_theory="EuclideanDivision">
      <ip_library name="int"/>
      <ip_qualid name="Div_bound"/>
     </pr_pos>
-    <pr_pos name="Mod_bound" id="2213"
+    <pr_pos name="Mod_bound" id="2115"
      ip_theory="EuclideanDivision">
      <ip_library name="int"/>
      <ip_qualid name="Mod_bound"/>
     </pr_pos>
-    <pr_pos name="Mod_1" id="2218"
+    <pr_pos name="Mod_1" id="2120"
      ip_theory="EuclideanDivision">
      <ip_library name="int"/>
      <ip_qualid name="Mod_1"/>
     </pr_pos>
-    <pr_pos name="Div_1" id="2221"
+    <pr_pos name="Div_1" id="2123"
      ip_theory="EuclideanDivision">
      <ip_library name="int"/>
      <ip_qualid name="Div_1"/>
     </pr_pos>
-    <pr_pos name="Div_inf" id="2224"
+    <pr_pos name="Div_inf" id="2126"
      ip_theory="EuclideanDivision">
      <ip_library name="int"/>
      <ip_qualid name="Div_inf"/>
     </pr_pos>
-    <pr_pos name="Div_inf_neg" id="2229"
+    <pr_pos name="Div_inf_neg" id="2131"
      ip_theory="EuclideanDivision">
      <ip_library name="int"/>
      <ip_qualid name="Div_inf_neg"/>
     </pr_pos>
-    <pr_pos name="Mod_0" id="2234"
+    <pr_pos name="Mod_0" id="2136"
      ip_theory="EuclideanDivision">
      <ip_library name="int"/>
      <ip_qualid name="Mod_0"/>
     </pr_pos>
-    <pr_pos name="Div_1_left" id="2237"
+    <pr_pos name="Div_1_left" id="2139"
      ip_theory="EuclideanDivision">
      <ip_library name="int"/>
      <ip_qualid name="Div_1_left"/>
     </pr_pos>
-    <pr_pos name="Div_minus1_left" id="2240"
+    <pr_pos name="Div_minus1_left" id="2142"
      ip_theory="EuclideanDivision">
      <ip_library name="int"/>
      <ip_qualid name="Div_minus1_left"/>
     </pr_pos>
-    <pr_pos name="Mod_1_left" id="2243"
+    <pr_pos name="Mod_1_left" id="2145"
      ip_theory="EuclideanDivision">
      <ip_library name="int"/>
      <ip_qualid name="Mod_1_left"/>
     </pr_pos>
-    <pr_pos name="Mod_minus1_left" id="2246"
+    <pr_pos name="Mod_minus1_left" id="2148"
      ip_theory="EuclideanDivision">
      <ip_library name="int"/>
      <ip_qualid name="Mod_minus1_left"/>
     </pr_pos>
-    <pr_pos name="Div_mult" id="2249"
+    <pr_pos name="Div_mult" id="2151"
      ip_theory="EuclideanDivision">
      <ip_library name="int"/>
      <ip_qualid name="Div_mult"/>
     </pr_pos>
-    <pr_pos name="Mod_mult" id="2256"
+    <pr_pos name="Mod_mult" id="2158"
      ip_theory="EuclideanDivision">
      <ip_library name="int"/>
      <ip_qualid name="Mod_mult"/>
     </pr_pos>
-    <pr_pos name="Length_nonnegative" id="2850"
+    <pr_pos name="Length_nonnegative" id="2752"
      ip_theory="Length">
      <ip_library name="list"/>
      <ip_qualid name="Length_nonnegative"/>
     </pr_pos>
-    <pr_pos name="Length_nil" id="2853"
+    <pr_pos name="Length_nil" id="2755"
      ip_theory="Length">
      <ip_library name="list"/>
      <ip_qualid name="Length_nil"/>
     </pr_pos>
-    <pr_pos name="Append_assoc" id="4390"
+    <pr_pos name="Append_assoc" id="4292"
      ip_theory="Append">
      <ip_library name="list"/>
      <ip_qualid name="Append_assoc"/>
     </pr_pos>
-    <pr_pos name="Append_length" id="4400"
+    <pr_pos name="Append_length" id="4302"
      ip_theory="Append">
      <ip_library name="list"/>
      <ip_qualid name="Append_length"/>
     </pr_pos>
-    <pr_pos name="mem_append" id="4405"
+    <pr_pos name="mem_append" id="4307"
      ip_theory="Append">
      <ip_library name="list"/>
      <ip_qualid name="mem_append"/>
     </pr_pos>
-    <pr_pos name="mem_decomp" id="4412"
+    <pr_pos name="mem_decomp" id="4314"
      ip_theory="Append">
      <ip_library name="list"/>
      <ip_qualid name="mem_decomp"/>
     </pr_pos>
-    <pr_pos name="Select_eq" id="5026"
+    <pr_pos name="Select_eq" id="4928"
      ip_theory="State">
      <ip_library name="state"/>
      <ip_qualid name="Select_eq"/>
     </pr_pos>
-    <pr_pos name="Select_neq" id="5032"
+    <pr_pos name="Select_neq" id="4934"
      ip_theory="State">
      <ip_library name="state"/>
      <ip_qualid name="Select_neq"/>
     </pr_pos>
-    <pr_pos name="Const" id="5039"
+    <pr_pos name="Const" id="4941"
      ip_theory="State">
      <ip_library name="state"/>
      <ip_qualid name="Const"/>
     </pr_pos>
-    <pr_pos name="inversion_beval_t" id="5157"
+    <pr_pos name="ceval_deterministic_aux" id="5151"
      ip_theory="Imp">
      <ip_library name="imp"/>
-     <ip_qualid name="inversion_beval_t"/>
+     <ip_qualid name="ceval_deterministic_aux"/>
     </pr_pos>
-    <pr_pos name="inversion_beval_f" id="5164"
-     ip_theory="Imp">
-     <ip_library name="imp"/>
-     <ip_qualid name="inversion_beval_f"/>
-    </pr_pos>
-    <pr_pos name="ceval_deterministic" id="5263"
+    <pr_pos name="ceval_deterministic" id="5160"
      ip_theory="Imp">
      <ip_library name="imp"/>
      <ip_qualid name="ceval_deterministic"/>
     </pr_pos>
-    <pr_pos name="codeseq_at_app_right" id="5476"
+    <pr_pos name="codeseq_at_app_right" id="5259"
      ip_theory="Vm">
      <ip_library name="vm"/>
      <ip_qualid name="codeseq_at_app_right"/>
     </pr_pos>
-    <pr_pos name="codeseq_at_app_left" id="5485"
+    <pr_pos name="codeseq_at_app_left" id="5268"
      ip_theory="Vm">
      <ip_library name="vm"/>
      <ip_qualid name="codeseq_at_app_left"/>
     </pr_pos>
-    <pr_pos name="transition_star_def" id="5795"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="transition_star_def"/>
-    </pr_pos>
-    <pr_pos name="transZero" id="5804"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="transZero"/>
-    </pr_pos>
-    <pr_pos name="transOne" id="5809"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="transOne"/>
-    </pr_pos>
-    <pr_pos name="transition_star_one" id="5818"
+    <pr_pos name="transition_star_one" id="5564"
      ip_theory="Vm">
      <ip_library name="vm"/>
      <ip_qualid name="transition_star_one"/>
     </pr_pos>
-    <pr_pos name="trans_star" id="5825"
-     ip_theory="Vm">
-     <ip_library name="vm"/>
-     <ip_qualid name="trans_star"/>
-    </pr_pos>
-    <pr_pos name="transition_star_transitive" id="5836"
+    <pr_pos name="transition_star_transitive" id="5571"
      ip_theory="Vm">
      <ip_library name="vm"/>
      <ip_qualid name="transition_star_transitive"/>
     </pr_pos>
-    <pr_pos name="seq_wp_lemma" id="6017"
+    <pr_pos name="seq_wp_lemma" id="5752"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="seq_wp_lemma"/>
     </pr_pos>
-    <pr_pos name="fork_wp_lemma" id="6095"
+    <pr_pos name="fork_wp_lemma" id="5830"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="fork_wp_lemma"/>
     </pr_pos>
-    <pr_pos name="towp_wp_lemma" id="6155"
+    <pr_pos name="towp_wp_lemma" id="5890"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="towp_wp_lemma"/>
     </pr_pos>
-    <pr_pos name="trivial_pre_lemma" id="6228"
+    <pr_pos name="trivial_pre_lemma" id="5963"
      ip_theory="Compiler_logic">
      <ip_library name="logic"/>
      <ip_qualid name="trivial_pre_lemma"/>
     </pr_pos>
-    <pr_pos name="ibinop_pre_lemma" id="6497"
-     ip_theory="VM_arith_instr_spec">
-     <ip_library name="specs"/>
-     <ip_qualid name="ibinop_pre_lemma"/>
-    </pr_pos>
-    <pr_pos name="ibinop_post_lemma" id="6561"
-     ip_theory="VM_arith_instr_spec">
-     <ip_library name="specs"/>
-     <ip_qualid name="ibinop_post_lemma"/>
-    </pr_pos>
-    <pr_pos name="icjump_post_lemma" id="6848"
-     ip_theory="VM_bool_instr_spec">
-     <ip_library name="specs"/>
-     <ip_qualid name="icjump_post_lemma"/>
-    </pr_pos>
-    <pr_pos name="beq_lemma" id="6892"
-     ip_theory="VM_bool_instr_spec">
-     <ip_library name="specs"/>
-     <ip_qualid name="beq_lemma"/>
-    </pr_pos>
-    <pr_pos name="bne_lemma" id="6912"
-     ip_theory="VM_bool_instr_spec">
-     <ip_library name="specs"/>
-     <ip_qualid name="bne_lemma"/>
-    </pr_pos>
-    <pr_pos name="ble_lemma" id="6932"
-     ip_theory="VM_bool_instr_spec">
-     <ip_library name="specs"/>
-     <ip_qualid name="ble_lemma"/>
-    </pr_pos>
-    <pr_pos name="bgt_lemma" id="6952"
-     ip_theory="VM_bool_instr_spec">
-     <ip_library name="specs"/>
-     <ip_qualid name="bgt_lemma"/>
-    </pr_pos>
-    <pr_pos name="isetvar_pre_lemma" id="7093"
-     ip_theory="VM_bool_instr_spec">
-     <ip_library name="specs"/>
-     <ip_qualid name="isetvar_pre_lemma"/>
-    </pr_pos>
-    <pr_pos name="isetvar_post_lemma" id="7150"
-     ip_theory="VM_bool_instr_spec">
-     <ip_library name="specs"/>
-     <ip_qualid name="isetvar_post_lemma"/>
-    </pr_pos>
-    <pr_pos name="aexpr_post_lemma" id="7254"
-     ip_theory="Compile_aexpr">
-     <ip_qualid name="aexpr_post_lemma"/>
-    </pr_pos>
-    <pr_pos name="com_pre_lemma" id="10969"
-     ip_theory="Compile_com">
-     <ip_qualid name="com_pre_lemma"/>
-    </pr_pos>
-    <pr_pos name="com_post_lemma" id="11080"
-     ip_theory="Compile_com">
-     <ip_qualid name="com_post_lemma"/>
-    </pr_pos>
     <meta name="remove_logic">
      <meta_arg_ls id="10"/>
     </meta>
@@ -3204,424 +2559,328 @@
      <meta_arg_ls id="15"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="787"/>
+     <meta_arg_ls id="689"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="788"/>
+     <meta_arg_ls id="690"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="791"/>
+     <meta_arg_ls id="693"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="1957"/>
+     <meta_arg_ls id="1859"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="1958"/>
+     <meta_arg_ls id="1860"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="1959"/>
+     <meta_arg_ls id="1861"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="2027"/>
+     <meta_arg_ls id="1929"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="2072"/>
+     <meta_arg_ls id="1974"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="2197"/>
+     <meta_arg_ls id="2099"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="2200"/>
+     <meta_arg_ls id="2102"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="2857"/>
+     <meta_arg_ls id="2759"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="3746"/>
+     <meta_arg_ls id="3648"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="3755"/>
+     <meta_arg_ls id="3657"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="3769"/>
+     <meta_arg_ls id="3671"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5000"/>
+     <meta_arg_ls id="4902"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5001"/>
+     <meta_arg_ls id="4903"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5038"/>
+     <meta_arg_ls id="4940"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5494"/>
+     <meta_arg_ls id="5580"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5503"/>
+     <meta_arg_ls id="5600"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5508"/>
+     <meta_arg_ls id="5619"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5513"/>
+     <meta_arg_ls id="5648"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5518"/>
+     <meta_arg_ls id="5673"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5519"/>
+     <meta_arg_ls id="5695"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5520"/>
+     <meta_arg_ls id="5721"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5521"/>
+     <meta_arg_ls id="5799"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5526"/>
+     <meta_arg_ls id="5860"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5531"/>
+     <meta_arg_ls id="5934"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5536"/>
+     <meta_arg_ls id="5984"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5541"/>
+     <meta_arg_ls id="5998"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5547"/>
+     <meta_arg_ls id="6034"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5769"/>
+     <meta_arg_ls id="6084"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5794"/>
+     <meta_arg_ls id="6146"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5845"/>
+     <meta_arg_ls id="6209"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5865"/>
+     <meta_arg_ls id="6242"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5884"/>
+     <meta_arg_ls id="6328"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5913"/>
+     <meta_arg_ls id="6341"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="5938"/>
-    </meta>
-    <meta name="remove_logic">
-     <meta_arg_ls id="5960"/>
-    </meta>
-    <meta name="remove_logic">
-     <meta_arg_ls id="5986"/>
-    </meta>
-    <meta name="remove_logic">
-     <meta_arg_ls id="6064"/>
-    </meta>
-    <meta name="remove_logic">
-     <meta_arg_ls id="6125"/>
-    </meta>
-    <meta name="remove_logic">
-     <meta_arg_ls id="6199"/>
-    </meta>
-    <meta name="remove_logic">
-     <meta_arg_ls id="6249"/>
-    </meta>
-    <meta name="remove_logic">
-     <meta_arg_ls id="6263"/>
-    </meta>
-    <meta name="remove_logic">
-     <meta_arg_ls id="6299"/>
-    </meta>
-    <meta name="remove_logic">
-     <meta_arg_ls id="6339"/>
+     <meta_arg_ls id="6354"/>
     </meta>
     <meta name="remove_logic">
      <meta_arg_ls id="6401"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="6464"/>
-    </meta>
-    <meta name="remove_logic">
-     <meta_arg_ls id="6519"/>
-    </meta>
-    <meta name="remove_logic">
-     <meta_arg_ls id="6633"/>
-    </meta>
-    <meta name="remove_logic">
-     <meta_arg_ls id="6646"/>
-    </meta>
-    <meta name="remove_logic">
-     <meta_arg_ls id="6659"/>
-    </meta>
-    <meta name="remove_logic">
-     <meta_arg_ls id="6706"/>
+     <meta_arg_ls id="6433"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="6738"/>
+     <meta_arg_ls id="6496"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="6801"/>
+     <meta_arg_ls id="6543"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="6879"/>
+     <meta_arg_ls id="6556"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="6899"/>
+     <meta_arg_ls id="6569"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="6919"/>
+     <meta_arg_ls id="6582"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="6939"/>
+     <meta_arg_ls id="6700"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="7064"/>
+     <meta_arg_ls id="6729"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="7112"/>
+     <meta_arg_ls id="6797"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="7205"/>
+     <meta_arg_ls id="8537"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="8980"/>
+     <meta_arg_ls id="8596"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="9039"/>
+     <meta_arg_ls id="10482"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="10925"/>
+     <meta_arg_ls id="10526"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="11001"/>
+     <meta_arg_ls id="10605"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="11141"/>
+     <meta_arg_ls id="10644"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="11180"/>
+     <meta_arg_ls id="10716"/>
     </meta>
     <meta name="remove_logic">
-     <meta_arg_ls id="11252"/>
-    </meta>
-    <meta name="remove_logic">
-     <meta_arg_ls id="11325"/>
-    </meta>
-    <meta name="remove_prop">
-     <meta_arg_pr id="1960"/>
-    </meta>
-    <meta name="remove_prop">
-     <meta_arg_pr id="1970"/>
-    </meta>
-    <meta name="remove_prop">
-     <meta_arg_pr id="1973"/>
-    </meta>
-    <meta name="remove_prop">
-     <meta_arg_pr id="1976"/>
-    </meta>
-    <meta name="remove_prop">
-     <meta_arg_pr id="1979"/>
-    </meta>
-    <meta name="remove_prop">
-     <meta_arg_pr id="1984"/>
-    </meta>
-    <meta name="remove_prop">
-     <meta_arg_pr id="1991"/>
-    </meta>
-    <meta name="remove_prop">
-     <meta_arg_pr id="1998"/>
+     <meta_arg_ls id="10789"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2016"/>
+     <meta_arg_pr id="1862"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2021"/>
+     <meta_arg_pr id="1872"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2024"/>
+     <meta_arg_pr id="1875"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2036"/>
+     <meta_arg_pr id="1878"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2039"/>
+     <meta_arg_pr id="1881"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2046"/>
+     <meta_arg_pr id="1886"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2051"/>
+     <meta_arg_pr id="1893"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2056"/>
+     <meta_arg_pr id="1900"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2057"/>
+     <meta_arg_pr id="1918"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2064"/>
+     <meta_arg_pr id="1923"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2077"/>
+     <meta_arg_pr id="1926"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2082"/>
+     <meta_arg_pr id="1938"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2203"/>
+     <meta_arg_pr id="1941"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2208"/>
+     <meta_arg_pr id="1948"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2213"/>
+     <meta_arg_pr id="1953"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2218"/>
+     <meta_arg_pr id="1958"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2221"/>
+     <meta_arg_pr id="1959"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2224"/>
+     <meta_arg_pr id="1966"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="2229"/>
-    </meta>
-    <meta name="remove_prop">
-     <meta_arg_pr id="2234"/>
-    </meta>
-    <meta name="remove_prop">
-     <meta_arg_pr id="2237"/>
-    </meta>
-    <meta name="remove_prop">
-     <meta_arg_pr id="2240"/>
-    </meta>
-    <meta name="remove_prop">
-     <meta_arg_pr id="2243"/>
-    </meta>
-    <meta name="remove_prop">
-     <meta_arg_pr id="2246"/>
-    </meta>
-    <meta name="remove_prop">
-     <meta_arg_pr id="2249"/>
-    </meta>
-    <meta name="remove_prop">
-     <meta_arg_pr id="2256"/>
-    </meta>
-    <meta name="remove_prop">
-     <meta_arg_pr id="2850"/>
-    </meta>
-    <meta name="remove_prop">
-     <meta_arg_pr id="2853"/>
+     <meta_arg_pr id="1979"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="4390"/>
+     <meta_arg_pr id="1984"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="4400"/>
+     <meta_arg_pr id="2105"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="4405"/>
+     <meta_arg_pr id="2110"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="4412"/>
+     <meta_arg_pr id="2115"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="5026"/>
+     <meta_arg_pr id="2120"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="5032"/>
+     <meta_arg_pr id="2123"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="5039"/>
+     <meta_arg_pr id="2126"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="5157"/>
+     <meta_arg_pr id="2131"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="5164"/>
+     <meta_arg_pr id="2136"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="5263"/>
+     <meta_arg_pr id="2139"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="5476"/>
+     <meta_arg_pr id="2142"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="5485"/>
+     <meta_arg_pr id="2145"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="5795"/>
+     <meta_arg_pr id="2148"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="5804"/>
+     <meta_arg_pr id="2151"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="5809"/>
+     <meta_arg_pr id="2158"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="5818"/>
+     <meta_arg_pr id="2752"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="5825"/>
+     <meta_arg_pr id="2755"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="5836"/>
+     <meta_arg_pr id="4292"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="6017"/>
+     <meta_arg_pr id="4302"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="6095"/>
+     <meta_arg_pr id="4307"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="6155"/>
+     <meta_arg_pr id="4314"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="6228"/>
+     <meta_arg_pr id="4928"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="6497"/>
+     <meta_arg_pr id="4934"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="6561"/>
+     <meta_arg_pr id="4941"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="6848"/>
+     <meta_arg_pr id="5151"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="6892"/>
+     <meta_arg_pr id="5160"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="6912"/>
+     <meta_arg_pr id="5259"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="6932"/>
+     <meta_arg_pr id="5268"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="6952"/>
+     <meta_arg_pr id="5564"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="7093"/>
+     <meta_arg_pr id="5571"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="7150"/>
+     <meta_arg_pr id="5752"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="7254"/>
+     <meta_arg_pr id="5830"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="10969"/>
+     <meta_arg_pr id="5890"/>
     </meta>
     <meta name="remove_prop">
-     <meta_arg_pr id="11080"/>
+     <meta_arg_pr id="5963"/>
     </meta>
     <meta name="remove_type">
      <meta_arg_ts id="2"/>
@@ -3639,45 +2898,42 @@
      <meta_arg_ts id="54"/>
     </meta>
     <meta name="remove_type">
-     <meta_arg_ts id="1502"/>
-    </meta>
-    <meta name="remove_type">
-     <meta_arg_ts id="5044"/>
+     <meta_arg_ts id="1404"/>
     </meta>
     <meta name="remove_type">
-     <meta_arg_ts id="5444"/>
+     <meta_arg_ts id="4946"/>
     </meta>
     <meta name="remove_type">
-     <meta_arg_ts id="5445"/>
+     <meta_arg_ts id="5227"/>
     </meta>
     <meta name="remove_type">
-     <meta_arg_ts id="5461"/>
+     <meta_arg_ts id="5228"/>
     </meta>
     <meta name="remove_type">
-     <meta_arg_ts id="5760"/>
+     <meta_arg_ts id="5244"/>
     </meta>
     <meta name="remove_type">
-     <meta_arg_ts id="5900"/>
+     <meta_arg_ts id="5635"/>
     </meta>
     <meta name="remove_type">
-     <meta_arg_ts id="5901"/>
+     <meta_arg_ts id="5636"/>
     </meta>
     <meta name="remove_type">
-     <meta_arg_ts id="5902"/>
+     <meta_arg_ts id="5637"/>
     </meta>
     <meta name="remove_type">
-     <meta_arg_ts id="5908"/>
+     <meta_arg_ts id="5643"/>
     </meta>
     <meta name="remove_type">
-     <meta_arg_ts id="6463"/>
+     <meta_arg_ts id="6208"/>
     </meta>
     <meta name="remove_type">
-     <meta_arg_ts id="6800"/>
+     <meta_arg_ts id="6495"/>
     </meta>
-    <goal name="WP_parameter compile_program.1.1" expl="1. postcondition" expanded="true">
-    <transf name="eliminate_builtin" expanded="true">
-     <goal name="WP_parameter compile_program.1.1.1" expl="1. postcondition" expanded="true">
-     <proof prover="0"><result status="valid" time="1.59"/></proof>
+    <goal name="WP_parameter compile_program.1.1" expl="1. postcondition">
+    <transf name="eliminate_builtin">
+     <goal name="WP_parameter compile_program.1.1.1" expl="1. postcondition">
+     <proof prover="3"><result status="valid" time="0.14"/></proof>
      </goal>
     </transf>
     </goal>
diff --git a/examples/in_progress/mini-compiler/compiler/why3shapes.gz b/examples/in_progress/mini-compiler/compiler/why3shapes.gz
index 0ee91601bd5de70b96a59ca5522fac4a0debfc00..54dce91649ac950245db41416c81f42029675920 100644
Binary files a/examples/in_progress/mini-compiler/compiler/why3shapes.gz and b/examples/in_progress/mini-compiler/compiler/why3shapes.gz differ
diff --git a/examples/in_progress/mini-compiler/imp.why b/examples/in_progress/mini-compiler/imp.why
index 3451e959e3cc865b31b40f1fafc6106708d00df4..0010f3a2df838392f18f78be17b7177cc51fa529 100644
--- a/examples/in_progress/mini-compiler/imp.why
+++ b/examples/in_progress/mini-compiler/imp.why
@@ -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
diff --git a/examples/in_progress/mini-compiler/imp/imp_Imp_ceval_deterministic_1.v b/examples/in_progress/mini-compiler/imp/imp_Imp_ceval_deterministic_1.v
deleted file mode 100644
index 019a153f09de6a3d3d2c052e3357c58e05a83fce..0000000000000000000000000000000000000000
--- a/examples/in_progress/mini-compiler/imp/imp_Imp_ceval_deterministic_1.v
+++ /dev/null
@@ -1,138 +0,0 @@
-(* 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.
-
diff --git a/examples/in_progress/mini-compiler/imp/why3session.xml b/examples/in_progress/mini-compiler/imp/why3session.xml
index 4415d89c688c72b411a63c3e7e47885875104c12..22dbce6d82f5c41f7715175baec51fddbc3a5ef9 100644
--- a/examples/in_progress/mini-compiler/imp/why3session.xml
+++ b/examples/in_progress/mini-compiler/imp/why3session.xml
@@ -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>
diff --git a/examples/in_progress/mini-compiler/imp/why3shapes.gz b/examples/in_progress/mini-compiler/imp/why3shapes.gz
index 1451148a4ecdb73e055cbd773b0107aa1752bf8d..a09967b2e899e00f83ce7dd75ea2e592f68ff113 100644
Binary files a/examples/in_progress/mini-compiler/imp/why3shapes.gz and b/examples/in_progress/mini-compiler/imp/why3shapes.gz differ
diff --git a/examples/in_progress/mini-compiler/logic.mlw b/examples/in_progress/mini-compiler/logic.mlw
index db7a085751ab76100160db7b125eee05f630245c..e24ef38cf806cd2af062bd8eaee456b2e87739d0 100644
--- a/examples/in_progress/mini-compiler/logic.mlw
+++ b/examples/in_progress/mini-compiler/logic.mlw
@@ -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
 
 
 
diff --git a/examples/in_progress/mini-compiler/logic/logic_Compiler_logic_WP_parameter_make_loop_hl_1.v b/examples/in_progress/mini-compiler/logic/logic_Compiler_logic_WP_parameter_make_loop_hl_1.v
deleted file mode 100644
index 7227e2a3c70972409ef05ba044092c765fb0a9f8..0000000000000000000000000000000000000000
--- a/examples/in_progress/mini-compiler/logic/logic_Compiler_logic_WP_parameter_make_loop_hl_1.v
+++ /dev/null
@@ -1,484 +0,0 @@
-(* 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.
-
diff --git a/examples/in_progress/mini-compiler/logic/why3session.xml b/examples/in_progress/mini-compiler/logic/why3session.xml
index 0a7aff60fd7ae00257bed79b397c5a93d85a497c..80ac097aa08b7c3d1041b6c9ccc2c1fbcf37f814 100644
--- a/examples/in_progress/mini-compiler/logic/why3session.xml
+++ b/examples/in_progress/mini-compiler/logic/why3session.xml
@@ -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>
diff --git a/examples/in_progress/mini-compiler/logic/why3shapes.gz b/examples/in_progress/mini-compiler/logic/why3shapes.gz
index 73279ef169162bbbe6a0502b9b022f7722b376e6..54acc8758326d16c5b5d6a6f5e317d09ba71254e 100644
Binary files a/examples/in_progress/mini-compiler/logic/why3shapes.gz and b/examples/in_progress/mini-compiler/logic/why3shapes.gz differ
diff --git a/examples/in_progress/mini-compiler/specs.mlw b/examples/in_progress/mini-compiler/specs.mlw
index 598884504aab69c88c4018d5978759c37c99e4bc..058988e996eff3bc4df3c99b91c9366b59c8a7e3 100644
--- a/examples/in_progress/mini-compiler/specs.mlw
+++ b/examples/in_progress/mini-compiler/specs.mlw
@@ -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  }
diff --git a/examples/in_progress/mini-compiler/specs/why3session.xml b/examples/in_progress/mini-compiler/specs/why3session.xml
index c0f0bda256c0b7c1b66f362b9c5a1dcb54480a5f..9e185927087aaac1baf2a6f7b81b6843e2056413 100644
--- a/examples/in_progress/mini-compiler/specs/why3session.xml
+++ b/examples/in_progress/mini-compiler/specs/why3session.xml
@@ -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">
diff --git a/examples/in_progress/mini-compiler/specs/why3shapes.gz b/examples/in_progress/mini-compiler/specs/why3shapes.gz
index 0c8cb1c1a403cdc9e6a54f580c06cee241dcedac..55c83cd301ce2c87d9fdd3439fd07d994bfe8579 100644
Binary files a/examples/in_progress/mini-compiler/specs/why3shapes.gz and b/examples/in_progress/mini-compiler/specs/why3shapes.gz differ
diff --git a/examples/in_progress/mini-compiler/vm.mlw b/examples/in_progress/mini-compiler/vm.mlw
index 34415c1b1e4eb239a0d712fc10c48dae98b98650..e945fe725aca16d5a3cdd3a0af9c9b1ae38719d5 100644
--- a/examples/in_progress/mini-compiler/vm.mlw
+++ b/examples/in_progress/mini-compiler/vm.mlw
@@ -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
 
 
diff --git a/examples/in_progress/mini-compiler/vm/why3session.xml b/examples/in_progress/mini-compiler/vm/why3session.xml
index 5aa16e17af0b9d18f136fe3d5cdc1cb9110e9756..00c9da59dfbe0ba932ad79799e52b4536e4a15c1 100644
--- a/examples/in_progress/mini-compiler/vm/why3session.xml
+++ b/examples/in_progress/mini-compiler/vm/why3session.xml
@@ -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>
diff --git a/examples/in_progress/mini-compiler/vm/why3shapes.gz b/examples/in_progress/mini-compiler/vm/why3shapes.gz
index bc7cfc5da687bb5c9c8e4be67be4d8d8841e3c35..7f5ecc8c7b78088b4587b6f974c699aca9247d13 100644
Binary files a/examples/in_progress/mini-compiler/vm/why3shapes.gz and b/examples/in_progress/mini-compiler/vm/why3shapes.gz differ