(*
 * Vericert: Verified high-level synthesis.
 * Copyright (C) 2020 Yann Herklotz <yann@yannherklotz.com>
 *
 * This program is free software: you can redistribute it and/or modify
 * it under the terms of the GNU General Public License as published by
 * the Free Software Foundation, either version 3 of the License, or
 * (at your option) any later version.
 *
 * This program is distributed in the hope that it will be useful,
 * but WITHOUT ANY WARRANTY; without even the implied warranty of
 * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE.  See the
 * GNU General Public License for more details.
 *
 * You should have received a copy of the GNU General Public License
 * along with this program.  If not, see <https://www.gnu.org/licenses/>.
 *)

From compcert Require RTL Op Maps Errors.
From compcert Require Import Maps Integers.
From vericert Require Import Vericertlib Verilog ValueInt HTL HTLgen AssocMap.
Require Import Lia.

Hint Resolve Maps.PTree.elements_keys_norepet : htlspec.
Hint Resolve Maps.PTree.elements_correct : htlspec.


forall (A B : Type) (f : mon A) (g : A -> mon B) (y : B) (s1 s3 : st) (i : st_incr s1 s3), (do X <- f; g X) s1 = OK y s3 i -> exists (x : A) (s2 : st) (i1 : st_prop s1 s2) (i2 : st_prop s2 s3), f s1 = OK x s2 i1 /\ g x s2 = OK y s3 i2

forall (A B : Type) (f : mon A) (g : A -> mon B) (y : B) (s1 s3 : st) (i : st_incr s1 s3), (do X <- f; g X) s1 = OK y s3 i -> exists (x : A) (s2 : st) (i1 : st_prop s1 s2) (i2 : st_prop s2 s3), f s1 = OK x s2 i1 /\ g x s2 = OK y s3 i2
A, B:Type
f:mon A
g:A -> mon B
y:B
s1, s3:st
i:st_incr s1 s3

(do X <- f; g X) s1 = OK y s3 i -> exists (x : A) (s2 : st) (i1 : st_prop s1 s2) (i2 : st_prop s2 s3), f s1 = OK x s2 i1 /\ g x s2 = OK y s3 i2
A, B:Type
f:mon A
g:A -> mon B
y:B
s1, s3:st
i:st_incr s1 s3

match f s1 with | Error msg => Error msg | OK a s' i => match g a s' with | Error msg => Error msg | OK b s'' i' => OK b s'' (st_trans s1 s' s'' i i') end end = OK y s3 i -> exists (x : A) (s2 : st) (i1 : st_prop s1 s2) (i2 : st_prop s2 s3), f s1 = OK x s2 i1 /\ g x s2 = OK y s3 i2
A, B:Type
f:mon A
g:A -> mon B
y:B
s1, s3:st
i:st_incr s1 s3
e:Errors.errmsg
H:Error e = OK y s3 i

exists (x : A) (s2 : st) (i1 : st_prop s1 s2) (i2 : st_prop s2 s3), Error e = OK x s2 i1 /\ g x s2 = OK y s3 i2
A, B:Type
f:mon A
g:A -> mon B
y:B
s1, s3:st
i:st_incr s1 s3
a:A
s':st
s:st_prop s1 s'
H:match g a s' with | Error msg => Error msg | OK b s'' i' => OK b s'' (st_trans s1 s' s'' s i') end = OK y s3 i
exists (x : A) (s2 : st) (i1 : st_prop s1 s2) (i2 : st_prop s2 s3), OK a s' s = OK x s2 i1 /\ g x s2 = OK y s3 i2
A, B:Type
f:mon A
g:A -> mon B
y:B
s1, s3:st
i:st_incr s1 s3
a:A
s':st
s:st_prop s1 s'
H:match g a s' with | Error msg => Error msg | OK b s'' i' => OK b s'' (st_trans s1 s' s'' s i') end = OK y s3 i

exists (x : A) (s2 : st) (i1 : st_prop s1 s2) (i2 : st_prop s2 s3), OK a s' s = OK x s2 i1 /\ g x s2 = OK y s3 i2
A, B:Type
f:mon A
g:A -> mon B
y:B
s1, s3:st
i:st_incr s1 s3
a:A
s':st
s:st_prop s1 s'
H:match g a s' with | Error msg => Error msg | OK b s'' i' => OK b s'' (st_trans s1 s' s'' s i') end = OK y s3 i

exists i2 : st_prop s' s3, OK a s' s = OK a s' s /\ g a s' = OK y s3 i2
A, B:Type
f:mon A
g:A -> mon B
y:B
s1, s3:st
i:st_incr s1 s3
a:A
s':st
s:st_prop s1 s'
s0:st_prop s' s3

exists i2 : st_prop s' s3, OK a s' s = OK a s' s /\ OK y s3 s0 = OK y s3 i2
exists s0; auto. Qed.

forall (A B C : Type) (f : mon (A * B)) (g : A -> B -> mon C) (z : C) (s1 s3 : st) (i : st_incr s1 s3), (do (X, Y)<- f; g X Y) s1 = OK z s3 i -> exists (x : A) (y : B) (s2 : st) (i1 : st_prop s1 s2) (i2 : st_prop s2 s3), f s1 = OK (x, y) s2 i1 /\ g x y s2 = OK z s3 i2

forall (A B C : Type) (f : mon (A * B)) (g : A -> B -> mon C) (z : C) (s1 s3 : st) (i : st_incr s1 s3), (do (X, Y)<- f; g X Y) s1 = OK z s3 i -> exists (x : A) (y : B) (s2 : st) (i1 : st_prop s1 s2) (i2 : st_prop s2 s3), f s1 = OK (x, y) s2 i1 /\ g x y s2 = OK z s3 i2
A, B, C:Type
f:mon (A * B)
g:A -> B -> mon C
z:C
s1, s3:st
i:st_incr s1 s3
H:(do xy <- f; g (fst xy) (snd xy)) s1 = OK z s3 i

exists (x : A) (y : B) (s2 : st) (i1 : st_prop s1 s2) (i2 : st_prop s2 s3), f s1 = OK (x, y) s2 i1 /\ g x y s2 = OK z s3 i2
A, B, C:Type
f:mon (A * B)
g:A -> B -> mon C
z:C
s1, s3:st
i:st_incr s1 s3
H:(do xy <- f; g (fst xy) (snd xy)) s1 = OK z s3 i

(exists (x : A * B) (s2 : st) (i1 : st_prop s1 s2) (i2 : st_prop s2 s3), f s1 = OK x s2 i1 /\ (fun xy : A * B => g (fst xy) (snd xy)) x s2 = OK z s3 i2) -> exists (x : A) (y : B) (s2 : st) (i1 : st_prop s1 s2) (i2 : st_prop s2 s3), f s1 = OK (x, y) s2 i1 /\ g x y s2 = OK z s3 i2
A, B, C:Type
f:mon (A * B)
g:A -> B -> mon C
z:C
s1, s3:st
i:st_incr s1 s3
H:(do xy <- f; g (fst xy) (snd xy)) s1 = OK z s3 i
x:A
y:B
s2:st
i1:st_prop s1 s2
i2:st_prop s2 s3
P:f s1 = OK (x, y) s2 i1
Q:g (fst (x, y)) (snd (x, y)) s2 = OK z s3 i2

exists (x : A) (y : B) (s2 : st) (i1 : st_prop s1 s2) (i2 : st_prop s2 s3), f s1 = OK (x, y) s2 i1 /\ g x y s2 = OK z s3 i2
A, B, C:Type
f:mon (A * B)
g:A -> B -> mon C
z:C
s1, s3:st
i:st_incr s1 s3
H:(do xy <- f; g (fst xy) (snd xy)) s1 = OK z s3 i
x:A
y:B
s2:st
i1:st_prop s1 s2
i2:st_prop s2 s3
P:f s1 = OK (x, y) s2 i1
Q:g x y s2 = OK z s3 i2

exists (x : A) (y : B) (s2 : st) (i1 : st_prop s1 s2) (i2 : st_prop s2 s3), f s1 = OK (x, y) s2 i1 /\ g x y s2 = OK z s3 i2
exists x; exists y; exists s2; exists i1; exists i2; auto. Qed. Ltac monadInv1 H := match type of H with | (OK _ _ _ = OK _ _ _) => inversion H; clear H; try subst | (Error _ _ = OK _ _ _) => discriminate | (ret _ _ = OK _ _ _) => inversion H; clear H; try subst | (error _ _ = OK _ _ _) => discriminate | (bind ?F ?G ?S = OK ?X ?S' ?I) => let x := fresh "x" in ( let s := fresh "s" in ( let i1 := fresh "INCR" in ( let i2 := fresh "INCR" in ( let EQ1 := fresh "EQ" in ( let EQ2 := fresh "EQ" in ( destruct (bind_inversion _ _ F G X S S' I H) as [x [s [i1 [i2 [EQ1 EQ2]]]]]; clear H; try (monadInv1 EQ2))))))) | (bind2 ?F ?G ?S = OK ?X ?S' ?I) => let x1 := fresh "x" in ( let x2 := fresh "x" in ( let s := fresh "s" in ( let i1 := fresh "INCR" in ( let i2 := fresh "INCR" in ( let EQ1 := fresh "EQ" in ( let EQ2 := fresh "EQ" in ( destruct (bind2_inversion _ _ _ F G X S S' I H) as [x1 [x2 [s [i1 [i2 [EQ1 EQ2]]]]]]; clear H; try (monadInv1 EQ2)))))))) end. Ltac monadInv H := match type of H with | (ret _ _ = OK _ _ _) => monadInv1 H | (error _ _ = OK _ _ _) => monadInv1 H | (bind ?F ?G ?S = OK ?X ?S' ?I) => monadInv1 H | (bind2 ?F ?G ?S = OK ?X ?S' ?I) => monadInv1 H | (?F _ _ _ _ _ _ _ _ = OK _ _ _) => ((progress simpl in H) || unfold F in H); monadInv1 H | (?F _ _ _ _ _ _ _ = OK _ _ _) => ((progress simpl in H) || unfold F in H); monadInv1 H | (?F _ _ _ _ _ _ = OK _ _ _) => ((progress simpl in H) || unfold F in H); monadInv1 H | (?F _ _ _ _ _ = OK _ _ _) => ((progress simpl in H) || unfold F in H); monadInv1 H | (?F _ _ _ _ = OK _ _ _) => ((progress simpl in H) || unfold F in H); monadInv1 H | (?F _ _ _ = OK _ _ _) => ((progress simpl in H) || unfold F in H); monadInv1 H | (?F _ _ = OK _ _ _) => ((progress simpl in H) || unfold F in H); monadInv1 H | (?F _ = OK _ _ _) => ((progress simpl in H) || unfold F in H); monadInv1 H end. (** * Relational specification of the translation *) (** We now define inductive predicates that characterise the fact that the statemachine that is created by the translation contains the correct translations for each of the elements *) Inductive tr_instr (fin rtrn st stk : reg) : RTL.instruction -> stmnt -> stmnt -> Prop := | tr_instr_Inop : forall n, Z.pos n <= Int.max_unsigned -> tr_instr fin rtrn st stk (RTL.Inop n) Vskip (state_goto st n) | tr_instr_Iop : forall n op args dst s s' e i, Z.pos n <= Int.max_unsigned -> translate_instr op args s = OK e s' i -> tr_instr fin rtrn st stk (RTL.Iop op args dst n) (Vnonblock (Vvar dst) e) (state_goto st n) | tr_instr_Icond : forall n1 n2 cond args s s' i c, Z.pos n1 <= Int.max_unsigned -> Z.pos n2 <= Int.max_unsigned -> translate_condition cond args s = OK c s' i -> tr_instr fin rtrn st stk (RTL.Icond cond args n1 n2) Vskip (state_cond st c n1 n2) | tr_instr_Ireturn_None : tr_instr fin rtrn st stk (RTL.Ireturn None) (Vseq (block fin (Vlit (ZToValue 1%Z))) (block rtrn (Vlit (ZToValue 0%Z)))) Vskip | tr_instr_Ireturn_Some : forall r, tr_instr fin rtrn st stk (RTL.Ireturn (Some r)) (Vseq (block fin (Vlit (ZToValue 1%Z))) (block rtrn (Vvar r))) Vskip | tr_instr_Iload : forall mem addr args s s' i c dst n, Z.pos n <= Int.max_unsigned -> translate_arr_access mem addr args stk s = OK c s' i -> tr_instr fin rtrn st stk (RTL.Iload mem addr args dst n) (nonblock dst c) (state_goto st n) | tr_instr_Istore : forall mem addr args s s' i c src n, Z.pos n <= Int.max_unsigned -> translate_arr_access mem addr args stk s = OK c s' i -> tr_instr fin rtrn st stk (RTL.Istore mem addr args src n) (Vnonblock c (Vvar src)) (state_goto st n).
tr_instr_Ijumptable :

forall cexpr tbl r, cexpr = tbl_to_case_expr st tbl -> tr_instr fin rtrn st stk (RTL.Ijumptable r tbl) (Vskip) (Vcase (Vvar r) cexpr (Some Vskip)).

Hint Constructors tr_instr : htlspec.

Inductive tr_code (c : RTL.code) (pc : RTL.node) (i : RTL.instruction) (stmnts trans : PTree.t stmnt)
          (fin rtrn st stk : reg) : Prop :=
  tr_code_intro :
    forall s t,
      c!pc = Some i ->
      stmnts!pc = Some s ->
      trans!pc = Some t ->
      tr_instr fin rtrn st stk i s t ->
      tr_code c pc i stmnts trans fin rtrn st stk.
Hint Constructors tr_code : htlspec.

Inductive tr_module (f : RTL.function) : module -> Prop :=
  tr_module_intro :
    forall data control fin rtrn st stk stk_len m start rst clk scldecls arrdecls wf,
      m = (mkmodule f.(RTL.fn_params)
                        data
                        control
                        f.(RTL.fn_entrypoint)
                        st stk stk_len fin rtrn start rst clk scldecls arrdecls wf) ->
      (forall pc i, Maps.PTree.get pc f.(RTL.fn_code) = Some i ->
               tr_code f.(RTL.fn_code) pc i data control fin rtrn st stk) ->
      stk_len = Z.to_nat (f.(RTL.fn_stacksize) / 4) ->
      Z.modulo (f.(RTL.fn_stacksize)) 4 = 0 ->
      0 <= f.(RTL.fn_stacksize) < Integers.Ptrofs.modulus ->
      st = ((RTL.max_reg_function f) + 1)%positive ->
      fin = ((RTL.max_reg_function f) + 2)%positive ->
      rtrn = ((RTL.max_reg_function f) + 3)%positive ->
      stk = ((RTL.max_reg_function f) + 4)%positive ->
      start = ((RTL.max_reg_function f) + 5)%positive ->
      rst = ((RTL.max_reg_function f) + 6)%positive ->
      clk = ((RTL.max_reg_function f) + 7)%positive ->
      tr_module f m.
Hint Constructors tr_module : htlspec.


forall (sz : nat) (s s' : st) (x : reg) (i : st_prop s s') (iop : option io), create_reg iop sz s = OK x s' i -> st_datapath s = st_datapath s'

forall (sz : nat) (s s' : st) (x : reg) (i : st_prop s s') (iop : option io), create_reg iop sz s = OK x s' i -> st_datapath s = st_datapath s'
sz:nat
s, s':st
x:reg
i:st_prop s s'
iop:option io
H:create_reg iop sz s = OK x s' i

st_datapath s = st_datapath s'
sz:nat
s:st
iop:option io
i:st_prop s {| st_st := st_st s; st_freshreg := Pos.succ (st_freshreg s); st_freshstate := st_freshstate s; st_scldecls := AssocMap.set (st_freshreg s) (iop, VScalar sz) (st_scldecls s); st_arrdecls := st_arrdecls s; st_datapath := st_datapath s; st_controllogic := st_controllogic s |}

st_datapath s = st_datapath {| st_st := st_st s; st_freshreg := Pos.succ (st_freshreg s); st_freshstate := st_freshstate s; st_scldecls := AssocMap.set (st_freshreg s) (iop, VScalar sz) (st_scldecls s); st_arrdecls := st_arrdecls s; st_datapath := st_datapath s; st_controllogic := st_controllogic s |}
trivial. Qed. Hint Resolve create_reg_datapath_trans : htlspec.

forall (sz : nat) (s s' : st) (x : reg) (i : st_prop s s') (iop : option io), create_reg iop sz s = OK x s' i -> st_controllogic s = st_controllogic s'

forall (sz : nat) (s s' : st) (x : reg) (i : st_prop s s') (iop : option io), create_reg iop sz s = OK x s' i -> st_controllogic s = st_controllogic s'
sz:nat
s, s':st
x:reg
i:st_prop s s'
iop:option io
H:create_reg iop sz s = OK x s' i

st_controllogic s = st_controllogic s'
sz:nat
s:st
iop:option io
i:st_prop s {| st_st := st_st s; st_freshreg := Pos.succ (st_freshreg s); st_freshstate := st_freshstate s; st_scldecls := AssocMap.set (st_freshreg s) (iop, VScalar sz) (st_scldecls s); st_arrdecls := st_arrdecls s; st_datapath := st_datapath s; st_controllogic := st_controllogic s |}

st_controllogic s = st_controllogic {| st_st := st_st s; st_freshreg := Pos.succ (st_freshreg s); st_freshstate := st_freshstate s; st_scldecls := AssocMap.set (st_freshreg s) (iop, VScalar sz) (st_scldecls s); st_arrdecls := st_arrdecls s; st_datapath := st_datapath s; st_controllogic := st_controllogic s |}
trivial. Qed. Hint Resolve create_reg_controllogic_trans : htlspec.

forall (sz : nat) (s s' : st) (x : unit) (i : st_prop s s') (iop : option io) (r : reg), declare_reg iop r sz s = OK x s' i -> st_datapath s = st_datapath s'

forall (sz : nat) (s s' : st) (x : unit) (i : st_prop s s') (iop : option io) (r : reg), declare_reg iop r sz s = OK x s' i -> st_datapath s = st_datapath s'
sz:nat
s, s':st
x:unit
i:st_prop s s'
iop:option io
r:reg
H:declare_reg iop r sz s = OK x s' i

st_datapath s = st_datapath s'
sz:nat
s:st
iop:option io
r:reg
i:st_prop s {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := AssocMap.set r (iop, VScalar sz) (st_scldecls s); st_arrdecls := st_arrdecls s; st_datapath := st_datapath s; st_controllogic := st_controllogic s |}

st_datapath s = st_datapath {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := AssocMap.set r (iop, VScalar sz) (st_scldecls s); st_arrdecls := st_arrdecls s; st_datapath := st_datapath s; st_controllogic := st_controllogic s |}
trivial. Qed. Hint Resolve create_reg_datapath_trans : htlspec.

forall (sz : nat) (s s' : st) (x : unit) (i : st_prop s s') (iop : option io) (r : reg), declare_reg iop r sz s = OK x s' i -> st_controllogic s = st_controllogic s'

forall (sz : nat) (s s' : st) (x : unit) (i : st_prop s s') (iop : option io) (r : reg), declare_reg iop r sz s = OK x s' i -> st_controllogic s = st_controllogic s'
sz:nat
s, s':st
x:unit
i:st_prop s s'
iop:option io
r:reg
H:declare_reg iop r sz s = OK x s' i

st_controllogic s = st_controllogic s'
sz:nat
s:st
iop:option io
r:reg
i:st_prop s {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := AssocMap.set r (iop, VScalar sz) (st_scldecls s); st_arrdecls := st_arrdecls s; st_datapath := st_datapath s; st_controllogic := st_controllogic s |}

st_controllogic s = st_controllogic {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := AssocMap.set r (iop, VScalar sz) (st_scldecls s); st_arrdecls := st_arrdecls s; st_datapath := st_datapath s; st_controllogic := st_controllogic s |}
trivial. Qed. Hint Resolve create_reg_controllogic_trans : htlspec.

forall (sz : nat) (s s' : st) (x : unit) (i : st_prop s s') (iop : option io) (r : reg), declare_reg iop r sz s = OK x s' i -> st_freshreg s = st_freshreg s'

forall (sz : nat) (s s' : st) (x : unit) (i : st_prop s s') (iop : option io) (r : reg), declare_reg iop r sz s = OK x s' i -> st_freshreg s = st_freshreg s'
inversion 1; auto. Qed. Hint Resolve declare_reg_freshreg_trans : htlspec.

forall (sz ln : nat) (s s' : st) (x : reg * nat) (i : st_prop s s') (iop : option io), create_arr iop sz ln s = OK x s' i -> st_datapath s = st_datapath s'

forall (sz ln : nat) (s s' : st) (x : reg * nat) (i : st_prop s s') (iop : option io), create_arr iop sz ln s = OK x s' i -> st_datapath s = st_datapath s'
sz, ln:nat
s, s':st
x:(reg * nat)%type
i:st_prop s s'
iop:option io
H:create_arr iop sz ln s = OK x s' i

st_datapath s = st_datapath s'
sz, ln:nat
s:st
iop:option io
i:st_prop s {| st_st := st_st s; st_freshreg := Pos.succ (st_freshreg s); st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := AssocMap.set (st_freshreg s) (iop, VArray sz ln) (st_arrdecls s); st_datapath := st_datapath s; st_controllogic := st_controllogic s |}

st_datapath s = st_datapath {| st_st := st_st s; st_freshreg := Pos.succ (st_freshreg s); st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := AssocMap.set (st_freshreg s) (iop, VArray sz ln) (st_arrdecls s); st_datapath := st_datapath s; st_controllogic := st_controllogic s |}
trivial. Qed. Hint Resolve create_arr_datapath_trans : htlspec.

forall (sz ln : nat) (s s' : st) (x : reg * nat) (i : st_prop s s') (iop : option io), create_arr iop sz ln s = OK x s' i -> st_controllogic s = st_controllogic s'

forall (sz ln : nat) (s s' : st) (x : reg * nat) (i : st_prop s s') (iop : option io), create_arr iop sz ln s = OK x s' i -> st_controllogic s = st_controllogic s'
sz, ln:nat
s, s':st
x:(reg * nat)%type
i:st_prop s s'
iop:option io
H:create_arr iop sz ln s = OK x s' i

st_controllogic s = st_controllogic s'
sz, ln:nat
s:st
iop:option io
i:st_prop s {| st_st := st_st s; st_freshreg := Pos.succ (st_freshreg s); st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := AssocMap.set (st_freshreg s) (iop, VArray sz ln) (st_arrdecls s); st_datapath := st_datapath s; st_controllogic := st_controllogic s |}

st_controllogic s = st_controllogic {| st_st := st_st s; st_freshreg := Pos.succ (st_freshreg s); st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := AssocMap.set (st_freshreg s) (iop, VArray sz ln) (st_arrdecls s); st_datapath := st_datapath s; st_controllogic := st_controllogic s |}
trivial. Qed. Hint Resolve create_arr_controllogic_trans : htlspec.

forall (s s' x : st) (i : st_prop s s'), get s = OK x s' i -> s = x

forall (s s' x : st) (i : st_prop s s'), get s = OK x s' i -> s = x
s, s', x:st
i:st_prop s s'
H:get s = OK x s' i
H1:s = x
H2:x = s'

s' = s'
trivial. Qed. Hint Resolve get_refl_x : htlspec.

forall (s s' x : st) (i : st_prop s s'), get s = OK x s' i -> s = s'

forall (s s' x : st) (i : st_prop s s'), get s = OK x s' i -> s = s'
s, s', x:st
i:st_prop s s'
H:get s = OK x s' i
H1:s = x
H2:x = s'

s' = s'
trivial. Qed. Hint Resolve get_refl_s : htlspec. Ltac inv_incr := repeat match goal with | [ H: create_reg _ _ ?s = OK _ ?s' _ |- _ ] => let H1 := fresh "H" in assert (H1 := H); eapply create_reg_datapath_trans in H; eapply create_reg_controllogic_trans in H1 | [ H: create_arr _ _ _ ?s = OK _ ?s' _ |- _ ] => let H1 := fresh "H" in assert (H1 := H); eapply create_arr_datapath_trans in H; eapply create_arr_controllogic_trans in H1 | [ H: get ?s = OK _ _ _ |- _ ] => let H1 := fresh "H" in assert (H1 := H); apply get_refl_x in H; apply get_refl_s in H1; subst | [ H: st_prop _ _ |- _ ] => unfold st_prop in H; destruct H | [ H: st_incr _ _ |- _ ] => destruct st_incr end.

forall (A : Type) (f : A -> forall x0 : state, res unit x0) (l : list A) (cs cs' : st) (ci : st_prop cs cs'), (forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_controllogic s = st_controllogic s') -> HTLMonadExtra.collectlist f l cs = OK tt cs' ci -> st_controllogic cs = st_controllogic cs'

forall (A : Type) (f : A -> forall x0 : state, res unit x0) (l : list A) (cs cs' : st) (ci : st_prop cs cs'), (forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_controllogic s = st_controllogic s') -> HTLMonadExtra.collectlist f l cs = OK tt cs' ci -> st_controllogic cs = st_controllogic cs'
A:Type
f:A -> forall x0 : state, res unit x0
cs':st
ci:st_prop cs' cs'
H:forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_controllogic s = st_controllogic s'

st_controllogic cs' = st_controllogic cs'
A:Type
f:A -> forall x0 : state, res unit x0
a:A
l:list A
IHl:forall (cs cs' : st) (ci : st_prop cs cs'), (forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_controllogic s = st_controllogic s') -> HTLMonadExtra.collectlist f l cs = OK tt cs' ci -> st_controllogic cs = st_controllogic cs'
cs, cs':st
ci:st_prop cs cs'
H:forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_controllogic s = st_controllogic s'
x:unit
s:st
INCR:st_prop cs s
INCR0:st_prop s cs'
EQ:f a cs = OK x s INCR
EQ0:HTLMonadExtra.collectlist f l s = OK tt cs' INCR0
st_controllogic cs = st_controllogic cs'
A:Type
f:A -> forall x0 : state, res unit x0
cs':st
ci:st_prop cs' cs'
H:forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_controllogic s = st_controllogic s'

st_controllogic cs' = st_controllogic cs'
trivial.
A:Type
f:A -> forall x0 : state, res unit x0
a:A
l:list A
IHl:forall (cs cs' : st) (ci : st_prop cs cs'), (forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_controllogic s = st_controllogic s') -> HTLMonadExtra.collectlist f l cs = OK tt cs' ci -> st_controllogic cs = st_controllogic cs'
cs, cs':st
ci:st_prop cs cs'
H:forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_controllogic s = st_controllogic s'
x:unit
s:st
INCR:st_prop cs s
INCR0:st_prop s cs'
EQ:f a cs = OK x s INCR
EQ0:HTLMonadExtra.collectlist f l s = OK tt cs' INCR0

st_controllogic cs = st_controllogic cs'
A:Type
f:A -> forall x0 : state, res unit x0
a:A
l:list A
IHl:forall (cs cs' : st) (ci : st_prop cs cs'), (forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_controllogic s = st_controllogic s') -> HTLMonadExtra.collectlist f l cs = OK tt cs' ci -> st_controllogic cs = st_controllogic cs'
cs, cs':st
ci:st_prop cs cs'
H:forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_controllogic s = st_controllogic s'
x:unit
s:st
INCR:st_prop cs s
INCR0:st_prop s cs'
EQ:st_controllogic cs = st_controllogic s
EQ0:HTLMonadExtra.collectlist f l s = OK tt cs' INCR0

st_controllogic cs = st_controllogic cs'
A:Type
f:A -> forall x0 : state, res unit x0
a:A
l:list A
IHl:forall (cs cs' : st) (ci : st_prop cs cs'), (forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_controllogic s = st_controllogic s') -> HTLMonadExtra.collectlist f l cs = OK tt cs' ci -> st_controllogic cs = st_controllogic cs'
cs, cs':st
ci:st_prop cs cs'
H:forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_controllogic s = st_controllogic s'
x:unit
s:st
INCR:st_prop cs s
INCR0:st_prop s cs'
EQ:st_controllogic cs = st_controllogic s
EQ0:HTLMonadExtra.collectlist f l s = OK tt cs' INCR0

st_controllogic s = st_controllogic cs'
eauto. Qed.

forall (A : Type) (f : A -> forall x0 : state, res unit x0) (l : list A) (cs cs' : st) (ci : st_prop cs cs'), (forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_datapath s = st_datapath s') -> HTLMonadExtra.collectlist f l cs = OK tt cs' ci -> st_datapath cs = st_datapath cs'

forall (A : Type) (f : A -> forall x0 : state, res unit x0) (l : list A) (cs cs' : st) (ci : st_prop cs cs'), (forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_datapath s = st_datapath s') -> HTLMonadExtra.collectlist f l cs = OK tt cs' ci -> st_datapath cs = st_datapath cs'
A:Type
f:A -> forall x0 : state, res unit x0
cs':st
ci:st_prop cs' cs'
H:forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_datapath s = st_datapath s'

st_datapath cs' = st_datapath cs'
A:Type
f:A -> forall x0 : state, res unit x0
a:A
l:list A
IHl:forall (cs cs' : st) (ci : st_prop cs cs'), (forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_datapath s = st_datapath s') -> HTLMonadExtra.collectlist f l cs = OK tt cs' ci -> st_datapath cs = st_datapath cs'
cs, cs':st
ci:st_prop cs cs'
H:forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_datapath s = st_datapath s'
x:unit
s:st
INCR:st_prop cs s
INCR0:st_prop s cs'
EQ:f a cs = OK x s INCR
EQ0:HTLMonadExtra.collectlist f l s = OK tt cs' INCR0
st_datapath cs = st_datapath cs'
A:Type
f:A -> forall x0 : state, res unit x0
cs':st
ci:st_prop cs' cs'
H:forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_datapath s = st_datapath s'

st_datapath cs' = st_datapath cs'
trivial.
A:Type
f:A -> forall x0 : state, res unit x0
a:A
l:list A
IHl:forall (cs cs' : st) (ci : st_prop cs cs'), (forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_datapath s = st_datapath s') -> HTLMonadExtra.collectlist f l cs = OK tt cs' ci -> st_datapath cs = st_datapath cs'
cs, cs':st
ci:st_prop cs cs'
H:forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_datapath s = st_datapath s'
x:unit
s:st
INCR:st_prop cs s
INCR0:st_prop s cs'
EQ:f a cs = OK x s INCR
EQ0:HTLMonadExtra.collectlist f l s = OK tt cs' INCR0

st_datapath cs = st_datapath cs'
A:Type
f:A -> forall x0 : state, res unit x0
a:A
l:list A
IHl:forall (cs cs' : st) (ci : st_prop cs cs'), (forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_datapath s = st_datapath s') -> HTLMonadExtra.collectlist f l cs = OK tt cs' ci -> st_datapath cs = st_datapath cs'
cs, cs':st
ci:st_prop cs cs'
H:forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_datapath s = st_datapath s'
x:unit
s:st
INCR:st_prop cs s
INCR0:st_prop s cs'
EQ:st_datapath cs = st_datapath s
EQ0:HTLMonadExtra.collectlist f l s = OK tt cs' INCR0

st_datapath cs = st_datapath cs'
A:Type
f:A -> forall x0 : state, res unit x0
a:A
l:list A
IHl:forall (cs cs' : st) (ci : st_prop cs cs'), (forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_datapath s = st_datapath s') -> HTLMonadExtra.collectlist f l cs = OK tt cs' ci -> st_datapath cs = st_datapath cs'
cs, cs':st
ci:st_prop cs cs'
H:forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_datapath s = st_datapath s'
x:unit
s:st
INCR:st_prop cs s
INCR0:st_prop s cs'
EQ:st_datapath cs = st_datapath s
EQ0:HTLMonadExtra.collectlist f l s = OK tt cs' INCR0

st_datapath s = st_datapath cs'
eauto. Qed.

forall (A : Type) (f : A -> forall x0 : state, res unit x0) (l : list A) (cs cs' : st) (ci : st_prop cs cs'), (forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_freshreg s = st_freshreg s') -> HTLMonadExtra.collectlist f l cs = OK tt cs' ci -> st_freshreg cs = st_freshreg cs'

forall (A : Type) (f : A -> forall x0 : state, res unit x0) (l : list A) (cs cs' : st) (ci : st_prop cs cs'), (forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_freshreg s = st_freshreg s') -> HTLMonadExtra.collectlist f l cs = OK tt cs' ci -> st_freshreg cs = st_freshreg cs'
A:Type
f:A -> forall x0 : state, res unit x0
cs':st
ci:st_prop cs' cs'
H:forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_freshreg s = st_freshreg s'

st_freshreg cs' = st_freshreg cs'
A:Type
f:A -> forall x0 : state, res unit x0
a:A
l:list A
IHl:forall (cs cs' : st) (ci : st_prop cs cs'), (forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_freshreg s = st_freshreg s') -> HTLMonadExtra.collectlist f l cs = OK tt cs' ci -> st_freshreg cs = st_freshreg cs'
cs, cs':st
ci:st_prop cs cs'
H:forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_freshreg s = st_freshreg s'
x:unit
s:st
INCR:st_prop cs s
INCR0:st_prop s cs'
EQ:f a cs = OK x s INCR
EQ0:HTLMonadExtra.collectlist f l s = OK tt cs' INCR0
st_freshreg cs = st_freshreg cs'
A:Type
f:A -> forall x0 : state, res unit x0
cs':st
ci:st_prop cs' cs'
H:forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_freshreg s = st_freshreg s'

st_freshreg cs' = st_freshreg cs'
trivial.
A:Type
f:A -> forall x0 : state, res unit x0
a:A
l:list A
IHl:forall (cs cs' : st) (ci : st_prop cs cs'), (forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_freshreg s = st_freshreg s') -> HTLMonadExtra.collectlist f l cs = OK tt cs' ci -> st_freshreg cs = st_freshreg cs'
cs, cs':st
ci:st_prop cs cs'
H:forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_freshreg s = st_freshreg s'
x:unit
s:st
INCR:st_prop cs s
INCR0:st_prop s cs'
EQ:f a cs = OK x s INCR
EQ0:HTLMonadExtra.collectlist f l s = OK tt cs' INCR0

st_freshreg cs = st_freshreg cs'
A:Type
f:A -> forall x0 : state, res unit x0
a:A
l:list A
IHl:forall (cs cs' : st) (ci : st_prop cs cs'), (forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_freshreg s = st_freshreg s') -> HTLMonadExtra.collectlist f l cs = OK tt cs' ci -> st_freshreg cs = st_freshreg cs'
cs, cs':st
ci:st_prop cs cs'
H:forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_freshreg s = st_freshreg s'
x:unit
s:st
INCR:st_prop cs s
INCR0:st_prop s cs'
EQ:st_freshreg cs = st_freshreg s
EQ0:HTLMonadExtra.collectlist f l s = OK tt cs' INCR0

st_freshreg cs = st_freshreg cs'
A:Type
f:A -> forall x0 : state, res unit x0
a:A
l:list A
IHl:forall (cs cs' : st) (ci : st_prop cs cs'), (forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_freshreg s = st_freshreg s') -> HTLMonadExtra.collectlist f l cs = OK tt cs' ci -> st_freshreg cs = st_freshreg cs'
cs, cs':st
ci:st_prop cs cs'
H:forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : A), f y s = OK x s' i -> st_freshreg s = st_freshreg s'
x:unit
s:st
INCR:st_prop cs s
INCR0:st_prop s cs'
EQ:st_freshreg cs = st_freshreg s
EQ0:HTLMonadExtra.collectlist f l s = OK tt cs' INCR0

st_freshreg s = st_freshreg cs'
eauto. Qed.

forall (io : option io) (n : nat) (l : list reg) (s s' : st) (i : st_prop s s'), HTLMonadExtra.collectlist (fun r : reg => declare_reg io r n) l s = OK tt s' i -> st_controllogic s = st_controllogic s'

forall (io : option io) (n : nat) (l : list reg) (s s' : st) (i : st_prop s s'), HTLMonadExtra.collectlist (fun r : reg => declare_reg io r n) l s = OK tt s' i -> st_controllogic s = st_controllogic s'
io:option Verilog.io
n:nat
l:list reg
s, s':st
i:st_prop s s'
H:HTLMonadExtra.collectlist (fun r : reg => declare_reg io r n) l s = OK tt s' i

st_controllogic s = st_controllogic s'
io:option Verilog.io
n:nat
l:list reg
s, s':st
i:st_prop s s'
H:HTLMonadExtra.collectlist (fun r : reg => declare_reg io r n) l s = OK tt s' i

forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : reg), (fun r : reg => declare_reg io r n) y s = OK x s' i -> st_controllogic s = st_controllogic s'
io:option Verilog.io
n:nat
l:list reg
s, s':st
i:st_prop s s'
H:HTLMonadExtra.collectlist (fun r : reg => declare_reg io r n) l s = OK tt s' i
s0:state
s'0:st
x:unit
i0:st_prop s0 s'0
y:reg
H0:(fun r : reg => declare_reg io r n) y s0 = OK x s'0 i0

st_controllogic s0 = st_controllogic s'0
eapply declare_reg_controllogic_trans. simpl in H0. eassumption. Qed.

forall (io : option io) (n : nat) (l : list reg) (s s' : st) (i : st_prop s s'), HTLMonadExtra.collectlist (fun r : reg => declare_reg io r n) l s = OK tt s' i -> st_datapath s = st_datapath s'

forall (io : option io) (n : nat) (l : list reg) (s s' : st) (i : st_prop s s'), HTLMonadExtra.collectlist (fun r : reg => declare_reg io r n) l s = OK tt s' i -> st_datapath s = st_datapath s'
io:option Verilog.io
n:nat
l:list reg
s, s':st
i:st_prop s s'
H:HTLMonadExtra.collectlist (fun r : reg => declare_reg io r n) l s = OK tt s' i

st_datapath s = st_datapath s'
io:option Verilog.io
n:nat
l:list reg
s, s':st
i:st_prop s s'
H:HTLMonadExtra.collectlist (fun r : reg => declare_reg io r n) l s = OK tt s' i

forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : reg), (fun r : reg => declare_reg io r n) y s = OK x s' i -> st_datapath s = st_datapath s'
io:option Verilog.io
n:nat
l:list reg
s, s':st
i:st_prop s s'
H:HTLMonadExtra.collectlist (fun r : reg => declare_reg io r n) l s = OK tt s' i
s0:state
s'0:st
x:unit
i0:st_prop s0 s'0
y:reg
H0:(fun r : reg => declare_reg io r n) y s0 = OK x s'0 i0

st_datapath s0 = st_datapath s'0
eapply declare_reg_datapath_trans. simpl in H0. eassumption. Qed.

forall (io : option io) (n : nat) (l : list reg) (s s' : st) (i : st_prop s s'), HTLMonadExtra.collectlist (fun r : reg => declare_reg io r n) l s = OK tt s' i -> st_freshreg s = st_freshreg s'

forall (io : option io) (n : nat) (l : list reg) (s s' : st) (i : st_prop s s'), HTLMonadExtra.collectlist (fun r : reg => declare_reg io r n) l s = OK tt s' i -> st_freshreg s = st_freshreg s'
io:option Verilog.io
n:nat
l:list reg
s, s':st
i:st_prop s s'
H:HTLMonadExtra.collectlist (fun r : reg => declare_reg io r n) l s = OK tt s' i

st_freshreg s = st_freshreg s'
io:option Verilog.io
n:nat
l:list reg
s, s':st
i:st_prop s s'
H:HTLMonadExtra.collectlist (fun r : reg => declare_reg io r n) l s = OK tt s' i

forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : reg), (fun r : reg => declare_reg io r n) y s = OK x s' i -> st_freshreg s = st_freshreg s'
io:option Verilog.io
n:nat
l:list reg
s, s':st
i:st_prop s s'
H:HTLMonadExtra.collectlist (fun r : reg => declare_reg io r n) l s = OK tt s' i
s0:state
s'0:st
x:unit
i0:st_prop s0 s'0
y:reg
H0:(fun r : reg => declare_reg io r n) y s0 = OK x s'0 i0
H2:tt = x
H3:{| st_st := st_st s0; st_freshreg := st_freshreg s0; st_freshstate := st_freshstate s0; st_scldecls := AssocMap.set y (io, VScalar n) (st_scldecls s0); st_arrdecls := st_arrdecls s0; st_datapath := st_datapath s0; st_controllogic := st_controllogic s0 |} = s'0

st_freshreg s0 = st_freshreg {| st_st := st_st s0; st_freshreg := st_freshreg s0; st_freshstate := st_freshstate s0; st_scldecls := AssocMap.set y (io, VScalar n) (st_scldecls s0); st_arrdecls := st_arrdecls s0; st_datapath := st_datapath s0; st_controllogic := st_controllogic s0 |}
auto. Qed. Ltac unfold_match H := match type of H with | context[match ?g with _ => _ end] => destruct g eqn:?; try discriminate end.

forall (op : Op.addressing) (args : list reg) (s : st) (r : expr) (s' : st) (i : st_prop s s'), translate_eff_addressing op args s = OK r s' i -> st_freshreg s = st_freshreg s'

forall (op : Op.addressing) (args : list reg) (s : st) (r : expr) (s' : st) (i : st_prop s s'), translate_eff_addressing op args s = OK r s' i -> st_freshreg s = st_freshreg s'
destruct op; intros; simpl in *; repeat (unfold_match H); inv H; auto. Qed. Hint Resolve translate_eff_addressing_freshreg_trans : htlspec.

forall (op : comparison) (args : list reg) (s : st) (r : expr) (s' : st) (i : st_prop s s'), translate_comparison op args s = OK r s' i -> st_freshreg s = st_freshreg s'

forall (op : comparison) (args : list reg) (s : st) (r : expr) (s' : st) (i : st_prop s s'), translate_comparison op args s = OK r s' i -> st_freshreg s = st_freshreg s'
destruct op; intros; simpl in *; repeat (unfold_match H); inv H; auto. Qed. Hint Resolve translate_comparison_freshreg_trans : htlspec.

forall (op : comparison) (args : list reg) (s : st) (r : expr) (s' : st) (i : st_prop s s'), translate_comparisonu op args s = OK r s' i -> st_freshreg s = st_freshreg s'

forall (op : comparison) (args : list reg) (s : st) (r : expr) (s' : st) (i : st_prop s s'), translate_comparisonu op args s = OK r s' i -> st_freshreg s = st_freshreg s'
destruct op; intros; simpl in *; repeat (unfold_match H); inv H; auto. Qed. Hint Resolve translate_comparisonu_freshreg_trans : htlspec.

forall (op : comparison) (args : list reg) (s : st) (r : expr) (s' : st) (i : st_prop s s') (n : int), translate_comparison_imm op args n s = OK r s' i -> st_freshreg s = st_freshreg s'

forall (op : comparison) (args : list reg) (s : st) (r : expr) (s' : st) (i : st_prop s s') (n : int), translate_comparison_imm op args n s = OK r s' i -> st_freshreg s = st_freshreg s'
destruct op; intros; simpl in *; repeat (unfold_match H); inv H; auto. Qed. Hint Resolve translate_comparison_imm_freshreg_trans : htlspec.

forall (op : comparison) (args : list reg) (s : st) (r : expr) (s' : st) (i : st_prop s s') (n : int), translate_comparison_immu op args n s = OK r s' i -> st_freshreg s = st_freshreg s'

forall (op : comparison) (args : list reg) (s : st) (r : expr) (s' : st) (i : st_prop s s') (n : int), translate_comparison_immu op args n s = OK r s' i -> st_freshreg s = st_freshreg s'
destruct op; intros; simpl in *; repeat (unfold_match H); inv H; auto. Qed. Hint Resolve translate_comparison_immu_freshreg_trans : htlspec.

forall (op : Op.condition) (args : list reg) (s : st) (r : expr) (s' : st) (i : st_prop s s'), translate_condition op args s = OK r s' i -> st_freshreg s = st_freshreg s'

forall (op : Op.condition) (args : list reg) (s : st) (r : expr) (s' : st) (i : st_prop s s'), translate_condition op args s = OK r s' i -> st_freshreg s = st_freshreg s'
destruct op; intros; simpl in *; repeat (unfold_match H); inv H; eauto with htlspec. Qed. Hint Resolve translate_condition_freshreg_trans : htlspec.

forall (op : Op.operation) (args : list reg) (s : st) (r : expr) (s' : st) (i : st_prop s s'), translate_instr op args s = OK r s' i -> st_freshreg s = st_freshreg s'

forall (op : Op.operation) (args : list reg) (s : st) (r : expr) (s' : st) (i : st_prop s s'), translate_instr op args s = OK r s' i -> st_freshreg s = st_freshreg s'
c:Op.condition
s:st
r:expr
s':st
i:st_prop s s'
r0, r1:reg
l0:list reg
H1:(do tc <- translate_condition c l0; ret (Vternary tc (Vvar r0) (Vvar r1))) s = OK r s' i

st_freshreg s = st_freshreg s'
c:Op.condition
s, s':st
i:st_prop s s'
r0, r1:reg
l0:list reg
x:expr
INCR:st_prop s s'
EQ:translate_condition c l0 s = OK x s' INCR
INCR0:st_prop s' s'

st_freshreg s = st_freshreg s'
eauto with htlspec. Qed. Hint Resolve translate_instr_freshreg_trans : htlspec.

forall (mem : AST.memory_chunk) (addr : Op.addressing) (args : list reg) (st0 : reg) (s : st) (r : expr) (s' : st) (i : st_prop s s'), translate_arr_access mem addr args st0 s = OK r s' i -> st_freshreg s = st_freshreg s'

forall (mem : AST.memory_chunk) (addr : Op.addressing) (args : list reg) (st0 : reg) (s : st) (r : expr) (s' : st) (i : st_prop s s'), translate_arr_access mem addr args st0 s = OK r s' i -> st_freshreg s = st_freshreg s'
mem:AST.memory_chunk
addr:Op.addressing
args:list reg
st0:reg
s:st
r:expr
s':st
i:st_prop s s'
H:translate_arr_access mem addr args st0 s = OK r s' i

st_freshreg s = st_freshreg s'
mem:AST.memory_chunk
addr:Op.addressing
args:list reg
st0:reg
s:st
r:expr
s':st
i:st_prop s s'
H:match mem with | AST.Mint32 => match addr with | Op.Aindexed off => match args with | nil => error (Errors.msg "HTLgen: translate_arr_access unsuported addressing") | r1 :: nil => if check_address_parameter_signed off then ret (Vvari st0 (Vbinop Vdivu (boplitz Vadd r1 off) (Vlit (ZToValue 4)))) else error (Errors.msg "HTLgen: translate_arr_access address out of bounds") | r1 :: _ :: _ => error (Errors.msg "HTLgen: translate_arr_access unsuported addressing") end | Op.Aindexed2scaled scale offset => match args with | nil => error (Errors.msg "HTLgen: translate_arr_access unsuported addressing") | r1 :: nil => error (Errors.msg "HTLgen: translate_arr_access unsuported addressing") | r1 :: r2 :: nil => if check_address_parameter_signed scale && check_address_parameter_signed offset then ret (Vvari st0 (Vbinop Vdivu (Vbinop Vadd (boplitz Vadd r1 offset) (boplitz Vmul r2 scale)) (Vlit (ZToValue 4)))) else error (Errors.msg "HTLgen: translate_arr_access address out of bounds") | r1 :: r2 :: _ :: _ => error (Errors.msg "HTLgen: translate_arr_access unsuported addressing") end | Op.Ainstack a => match args with | nil => if check_address_parameter_unsigned (Ptrofs.unsigned a) then ret (Vvari st0 (Vlit (ZToValue (Ptrofs.unsigned a / 4)))) else error (Errors.msg "HTLgen: eff_addressing out of bounds stack offset") | _ :: _ => error (Errors.msg "HTLgen: translate_arr_access unsuported addressing") end | _ => error (Errors.msg "HTLgen: translate_arr_access unsuported addressing") end | _ => error (Errors.msg "HTLgen: translate_arr_access unsuported addressing") end s = OK r s' i

st_freshreg s = st_freshreg s'
repeat (unfold_match H); inv H; eauto with htlspec. Qed. Hint Resolve translate_arr_access_freshreg_trans : htlspec.

forall (n n' : node) (st0 : stmnt) (s : st) (r : unit) (s' : st) (i : st_prop s s'), add_instr n n' st0 s = OK r s' i -> st_freshreg s = st_freshreg s'

forall (n n' : node) (st0 : stmnt) (s : st) (r : unit) (s' : st) (i : st_prop s s'), add_instr n n' st0 s = OK r s' i -> st_freshreg s = st_freshreg s'
n, n':node
st0:stmnt
s:st
r:unit
s':st
i:st_prop s s'
H:add_instr n n' st0 s = OK r s' i

st_freshreg s = st_freshreg s'
n, n':node
st0:stmnt
s:st
r:unit
s':st
i:st_prop s s'
H:match check_empty_node_datapath s n with | left STM => match check_empty_node_controllogic s n with | left TRANS => OK tt {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := st_arrdecls s; st_datapath := AssocMap.set n st0 (st_datapath s); st_controllogic := AssocMap.set n (state_goto (st_st s) n') (st_controllogic s) |} (add_instr_state_incr s n n' st0 STM TRANS) | right _ => Error (Errors.msg "HTL.add_instr") end | right _ => Error (Errors.msg "HTL.add_instr") end = OK r s' i

st_freshreg s = st_freshreg s'
n, n':node
st0:stmnt
s:st
r:unit
s':st
i:st_prop s s'
e:(st_datapath s) ! n = None
Heqs0:check_empty_node_datapath s n = left e
e0:(st_controllogic s) ! n = None
Heqs1:check_empty_node_controllogic s n = left e0
H:OK tt {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := st_arrdecls s; st_datapath := AssocMap.set n st0 (st_datapath s); st_controllogic := AssocMap.set n (state_goto (st_st s) n') (st_controllogic s) |} (add_instr_state_incr s n n' st0 e e0) = OK r s' i

st_freshreg s = st_freshreg s'
n, n':node
st0:stmnt
s:st
i:st_prop s {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := st_arrdecls s; st_datapath := AssocMap.set n st0 (st_datapath s); st_controllogic := AssocMap.set n (state_goto (st_st s) n') (st_controllogic s) |}
e:(st_datapath s) ! n = None
Heqs0:check_empty_node_datapath s n = left e
e0:(st_controllogic s) ! n = None
Heqs1:check_empty_node_controllogic s n = left e0

st_freshreg s = st_freshreg {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := st_arrdecls s; st_datapath := AssocMap.set n st0 (st_datapath s); st_controllogic := AssocMap.set n (state_goto (st_st s) n') (st_controllogic s) |}
auto. Qed. Hint Resolve add_instr_freshreg_trans : htlspec.

forall (n n0 n1 : node) (e : expr) (s : st) (r : unit) (s' : st) (i : st_prop s s'), add_branch_instr e n n0 n1 s = OK r s' i -> st_freshreg s = st_freshreg s'

forall (n n0 n1 : node) (e : expr) (s : st) (r : unit) (s' : st) (i : st_prop s s'), add_branch_instr e n n0 n1 s = OK r s' i -> st_freshreg s = st_freshreg s'
n, n0, n1:node
e:expr
s:st
r:unit
s':st
i:st_prop s s'
H:add_branch_instr e n n0 n1 s = OK r s' i

st_freshreg s = st_freshreg s'
n, n0, n1:node
e:expr
s:st
r:unit
s':st
i:st_prop s s'
H:match check_empty_node_datapath s n with | left NSTM => match check_empty_node_controllogic s n with | left NTRANS => OK tt {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := st_arrdecls s; st_datapath := AssocMap.set n Vskip (st_datapath s); st_controllogic := AssocMap.set n (state_cond (st_st s) e n0 n1) (st_controllogic s) |} (add_branch_instr_state_incr s e n n0 n1 NSTM NTRANS) | right _ => Error (Errors.msg "Htlgen: add_branch_instr") end | right _ => Error (Errors.msg "Htlgen: add_branch_instr") end = OK r s' i

st_freshreg s = st_freshreg s'
n, n0, n1:node
e:expr
s:st
r:unit
s':st
i:st_prop s s'
e0:(st_datapath s) ! n = None
Heqs0:check_empty_node_datapath s n = left e0
e1:(st_controllogic s) ! n = None
Heqs1:check_empty_node_controllogic s n = left e1
H:OK tt {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := st_arrdecls s; st_datapath := AssocMap.set n Vskip (st_datapath s); st_controllogic := AssocMap.set n (state_cond (st_st s) e n0 n1) (st_controllogic s) |} (add_branch_instr_state_incr s e n n0 n1 e0 e1) = OK r s' i

st_freshreg s = st_freshreg s'
n, n0, n1:node
e:expr
s:st
i:st_prop s {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := st_arrdecls s; st_datapath := AssocMap.set n Vskip (st_datapath s); st_controllogic := AssocMap.set n (state_cond (st_st s) e n0 n1) (st_controllogic s) |}
e0:(st_datapath s) ! n = None
Heqs0:check_empty_node_datapath s n = left e0
e1:(st_controllogic s) ! n = None
Heqs1:check_empty_node_controllogic s n = left e1

st_freshreg s = st_freshreg {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := st_arrdecls s; st_datapath := AssocMap.set n Vskip (st_datapath s); st_controllogic := AssocMap.set n (state_cond (st_st s) e n0 n1) (st_controllogic s) |}
auto. Qed. Hint Resolve add_branch_instr_freshreg_trans : htlspec.

forall (n1 : node) (n2 : stmnt) (s : st) (r : unit) (s' : st) (i : st_prop s s'), add_node_skip n1 n2 s = OK r s' i -> st_freshreg s = st_freshreg s'

forall (n1 : node) (n2 : stmnt) (s : st) (r : unit) (s' : st) (i : st_prop s s'), add_node_skip n1 n2 s = OK r s' i -> st_freshreg s = st_freshreg s'
n1:node
n2:stmnt
s:st
r:unit
s':st
i:st_prop s s'
H:add_node_skip n1 n2 s = OK r s' i

st_freshreg s = st_freshreg s'
n1:node
n2:stmnt
s:st
r:unit
s':st
i:st_prop s s'
H:match check_empty_node_datapath s n1 with | left STM => match check_empty_node_controllogic s n1 with | left TRANS => OK tt {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := st_arrdecls s; st_datapath := AssocMap.set n1 Vskip (st_datapath s); st_controllogic := AssocMap.set n1 n2 (st_controllogic s) |} (add_node_skip_state_incr s n1 n2 STM TRANS) | right _ => Error (Errors.msg "HTL.add_instr") end | right _ => Error (Errors.msg "HTL.add_instr") end = OK r s' i

st_freshreg s = st_freshreg s'
n1:node
n2:stmnt
s:st
r:unit
s':st
i:st_prop s s'
e:(st_datapath s) ! n1 = None
Heqs0:check_empty_node_datapath s n1 = left e
e0:(st_controllogic s) ! n1 = None
Heqs1:check_empty_node_controllogic s n1 = left e0
H:OK tt {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := st_arrdecls s; st_datapath := AssocMap.set n1 Vskip (st_datapath s); st_controllogic := AssocMap.set n1 n2 (st_controllogic s) |} (add_node_skip_state_incr s n1 n2 e e0) = OK r s' i

st_freshreg s = st_freshreg s'
n1:node
n2:stmnt
s:st
i:st_prop s {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := st_arrdecls s; st_datapath := AssocMap.set n1 Vskip (st_datapath s); st_controllogic := AssocMap.set n1 n2 (st_controllogic s) |}
e:(st_datapath s) ! n1 = None
Heqs0:check_empty_node_datapath s n1 = left e
e0:(st_controllogic s) ! n1 = None
Heqs1:check_empty_node_controllogic s n1 = left e0

st_freshreg s = st_freshreg {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := st_arrdecls s; st_datapath := AssocMap.set n1 Vskip (st_datapath s); st_controllogic := AssocMap.set n1 n2 (st_controllogic s) |}
auto. Qed. Hint Resolve add_node_skip_freshreg_trans : htlspec.

forall (n1 : node) (n2 : stmnt) (s : st) (r : unit) (s' : st) (i : st_prop s s'), add_instr_skip n1 n2 s = OK r s' i -> st_freshreg s = st_freshreg s'

forall (n1 : node) (n2 : stmnt) (s : st) (r : unit) (s' : st) (i : st_prop s s'), add_instr_skip n1 n2 s = OK r s' i -> st_freshreg s = st_freshreg s'
n1:node
n2:stmnt
s:st
r:unit
s':st
i:st_prop s s'
H:add_instr_skip n1 n2 s = OK r s' i

st_freshreg s = st_freshreg s'
n1:node
n2:stmnt
s:st
r:unit
s':st
i:st_prop s s'
H:match check_empty_node_datapath s n1 with | left STM => match check_empty_node_controllogic s n1 with | left TRANS => OK tt {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := st_arrdecls s; st_datapath := AssocMap.set n1 n2 (st_datapath s); st_controllogic := AssocMap.set n1 Vskip (st_controllogic s) |} (add_instr_skip_state_incr s n1 n2 STM TRANS) | right _ => Error (Errors.msg "HTL.add_instr") end | right _ => Error (Errors.msg "HTL.add_instr") end = OK r s' i

st_freshreg s = st_freshreg s'
n1:node
n2:stmnt
s:st
r:unit
s':st
i:st_prop s s'
e:(st_datapath s) ! n1 = None
Heqs0:check_empty_node_datapath s n1 = left e
e0:(st_controllogic s) ! n1 = None
Heqs1:check_empty_node_controllogic s n1 = left e0
H:OK tt {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := st_arrdecls s; st_datapath := AssocMap.set n1 n2 (st_datapath s); st_controllogic := AssocMap.set n1 Vskip (st_controllogic s) |} (add_instr_skip_state_incr s n1 n2 e e0) = OK r s' i

st_freshreg s = st_freshreg s'
n1:node
n2:stmnt
s:st
i:st_prop s {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := st_arrdecls s; st_datapath := AssocMap.set n1 n2 (st_datapath s); st_controllogic := AssocMap.set n1 Vskip (st_controllogic s) |}
e:(st_datapath s) ! n1 = None
Heqs0:check_empty_node_datapath s n1 = left e
e0:(st_controllogic s) ! n1 = None
Heqs1:check_empty_node_controllogic s n1 = left e0

st_freshreg s = st_freshreg {| st_st := st_st s; st_freshreg := st_freshreg s; st_freshstate := st_freshstate s; st_scldecls := st_scldecls s; st_arrdecls := st_arrdecls s; st_datapath := AssocMap.set n1 n2 (st_datapath s); st_controllogic := AssocMap.set n1 Vskip (st_controllogic s) |}
auto. Qed. Hint Resolve add_instr_skip_freshreg_trans : htlspec.

forall (fin ret st0 : reg) (instr : node * RTL.instruction) (s : st) (v : unit) (s' : st) (i : st_prop s s'), transf_instr fin ret st0 instr s = OK v s' i -> st_freshreg s = st_freshreg s'

forall (fin ret st0 : reg) (instr : node * RTL.instruction) (s : st) (v : unit) (s' : st) (i : st_prop s s'), transf_instr fin ret st0 instr s = OK v s' i -> st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
instr:(node * RTL.instruction)%type
s:st
v:unit
s':st
i:st_prop s s'
H:transf_instr fin ret0 st0 instr s = OK v s' i

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
instr:(node * RTL.instruction)%type
s:st
v:unit
s':st
i:st_prop s s'
n:node
i0:RTL.instruction
Heqp:instr = (n, i0)
H:transf_instr fin ret0 st0 (n, i0) s = OK v s' i

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
i0:RTL.instruction
H:transf_instr fin ret0 st0 (n, i0) s = OK v s' i

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
i0:RTL.instruction
H:match i0 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr n n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr n n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args st0; do _ <- declare_reg None dst 32; add_instr n n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args st0; add_instr n n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e n n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip n (Vseq (block fin (Vlit (ZToValue 1))) (block ret0 (Vvar r'))) | RTL.Ireturn None => add_instr_skip n (Vseq (block fin (Vlit (ZToValue 1))) (block ret0 (Vlit (ZToValue 0)))) end s = OK v s' i

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
o:Op.operation
l:list Registers.reg
r:Registers.reg
n0:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) = true
H:(do instr <- translate_instr o l; do _ <- declare_reg None r 32; add_instr n n0 (nonblock r instr)) s = OK v s' i

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
m:AST.memory_chunk
a:Op.addressing
l:list Registers.reg
r:Registers.reg
n0:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) = true
H:(do src <- translate_arr_access m a l st0; do _ <- declare_reg None r 32; add_instr n n0 (nonblock r src)) s = OK v s' i
st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
m:AST.memory_chunk
a:Op.addressing
l:list Registers.reg
r:Registers.reg
n0:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) = true
H:(do dst <- translate_arr_access m a l st0; add_instr n n0 (Vnonblock dst (Vvar r))) s = OK v s' i
st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
c:Op.condition
l:list Registers.reg
n0, n1:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) && (Z.pos n1 <=? Int.max_unsigned) = true
H:(do e <- translate_condition c l; add_branch_instr e n n0 n1) s = OK v s' i
st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
o:Op.operation
l:list Registers.reg
r:Registers.reg
n0:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) = true
H:(do instr <- translate_instr o l; do _ <- declare_reg None r 32; add_instr n n0 (nonblock r instr)) s = OK v s' i

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
o:Op.operation
l:list Registers.reg
r:Registers.reg
n0:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) = true
x:expr
s0:st
INCR:st_prop s s0
INCR0:st_prop s0 s'
EQ:translate_instr o l s = OK x s0 INCR
x0:unit
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:declare_reg None r 32 s0 = OK x0 s1 INCR1
EQ2:add_instr n n0 (nonblock r x) s1 = OK v s' INCR2

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
o:Op.operation
l:list Registers.reg
r:Registers.reg
n0:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) = true
x:expr
s0:st
INCR:st_prop s s0
INCR0:st_prop s0 s'
EQ:translate_instr o l s = OK x s0 INCR
x0:unit
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:declare_reg None r 32 s0 = OK x0 s1 INCR1
EQ2:st_freshreg s1 = st_freshreg s'

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
o:Op.operation
l:list Registers.reg
r:Registers.reg
n0:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) = true
x:expr
s0:st
INCR:st_prop s s0
INCR0:st_prop s0 s'
EQ:st_freshreg s = st_freshreg s0
x0:unit
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:declare_reg None r 32 s0 = OK x0 s1 INCR1
EQ2:st_freshreg s1 = st_freshreg s'

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
o:Op.operation
l:list Registers.reg
r:Registers.reg
n0:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) = true
x:expr
s0:st
INCR:st_prop s s0
INCR0:st_prop s0 s'
EQ:st_freshreg s = st_freshreg s0
x0:unit
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:st_freshreg s0 = st_freshreg s1
EQ2:st_freshreg s1 = st_freshreg s'

st_freshreg s = st_freshreg s'
congruence.
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
m:AST.memory_chunk
a:Op.addressing
l:list Registers.reg
r:Registers.reg
n0:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) = true
H:(do src <- translate_arr_access m a l st0; do _ <- declare_reg None r 32; add_instr n n0 (nonblock r src)) s = OK v s' i

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
m:AST.memory_chunk
a:Op.addressing
l:list Registers.reg
r:Registers.reg
n0:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) = true
x:expr
s0:st
INCR:st_prop s s0
INCR0:st_prop s0 s'
EQ:translate_arr_access m a l st0 s = OK x s0 INCR
x0:unit
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:declare_reg None r 32 s0 = OK x0 s1 INCR1
EQ2:add_instr n n0 (nonblock r x) s1 = OK v s' INCR2

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
m:AST.memory_chunk
a:Op.addressing
l:list Registers.reg
r:Registers.reg
n0:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) = true
x:expr
s0:st
INCR:st_prop s s0
INCR0:st_prop s0 s'
EQ:translate_arr_access m a l st0 s = OK x s0 INCR
x0:unit
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:declare_reg None r 32 s0 = OK x0 s1 INCR1
EQ2:st_freshreg s1 = st_freshreg s'

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
m:AST.memory_chunk
a:Op.addressing
l:list Registers.reg
r:Registers.reg
n0:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) = true
x:expr
s0:st
INCR:st_prop s s0
INCR0:st_prop s0 s'
EQ:st_freshreg s = st_freshreg s0
x0:unit
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:declare_reg None r 32 s0 = OK x0 s1 INCR1
EQ2:st_freshreg s1 = st_freshreg s'

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
m:AST.memory_chunk
a:Op.addressing
l:list Registers.reg
r:Registers.reg
n0:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) = true
x:expr
s0:st
INCR:st_prop s s0
INCR0:st_prop s0 s'
EQ:st_freshreg s = st_freshreg s0
x0:unit
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:st_freshreg s0 = st_freshreg s1
EQ2:st_freshreg s1 = st_freshreg s'

st_freshreg s = st_freshreg s'
congruence.
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
m:AST.memory_chunk
a:Op.addressing
l:list Registers.reg
r:Registers.reg
n0:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) = true
H:(do dst <- translate_arr_access m a l st0; add_instr n n0 (Vnonblock dst (Vvar r))) s = OK v s' i

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
m:AST.memory_chunk
a:Op.addressing
l:list Registers.reg
r:Registers.reg
n0:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) = true
x:expr
s0:st
INCR:st_prop s s0
INCR0:st_prop s0 s'
EQ:translate_arr_access m a l st0 s = OK x s0 INCR
EQ0:add_instr n n0 (Vnonblock x (Vvar r)) s0 = OK v s' INCR0

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
m:AST.memory_chunk
a:Op.addressing
l:list Registers.reg
r:Registers.reg
n0:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) = true
x:expr
s0:st
INCR:st_prop s s0
INCR0:st_prop s0 s'
EQ:translate_arr_access m a l st0 s = OK x s0 INCR
EQ0:st_freshreg s0 = st_freshreg s'

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
m:AST.memory_chunk
a:Op.addressing
l:list Registers.reg
r:Registers.reg
n0:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) = true
x:expr
s0:st
INCR:st_prop s s0
INCR0:st_prop s0 s'
EQ:st_freshreg s = st_freshreg s0
EQ0:st_freshreg s0 = st_freshreg s'

st_freshreg s = st_freshreg s'
congruence.
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
c:Op.condition
l:list Registers.reg
n0, n1:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) && (Z.pos n1 <=? Int.max_unsigned) = true
H:(do e <- translate_condition c l; add_branch_instr e n n0 n1) s = OK v s' i

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
c:Op.condition
l:list Registers.reg
n0, n1:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) && (Z.pos n1 <=? Int.max_unsigned) = true
x:expr
s0:st
INCR:st_prop s s0
INCR0:st_prop s0 s'
EQ:translate_condition c l s = OK x s0 INCR
EQ0:add_branch_instr x n n0 n1 s0 = OK v s' INCR0

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
c:Op.condition
l:list Registers.reg
n0, n1:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) && (Z.pos n1 <=? Int.max_unsigned) = true
x:expr
s0:st
INCR:st_prop s s0
INCR0:st_prop s0 s'
EQ:st_freshreg s = st_freshreg s0
EQ0:add_branch_instr x n n0 n1 s0 = OK v s' INCR0

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
s:st
v:unit
s':st
i:st_prop s s'
n:node
c:Op.condition
l:list Registers.reg
n0, n1:RTL.node
Heqb:(Z.pos n0 <=? Int.max_unsigned) && (Z.pos n1 <=? Int.max_unsigned) = true
x:expr
s0:st
INCR:st_prop s s0
INCR0:st_prop s0 s'
EQ:st_freshreg s = st_freshreg s0
EQ0:st_freshreg s0 = st_freshreg s'

st_freshreg s = st_freshreg s'
congruence. (*- inv EQ. apply add_node_skip_freshreg_trans in EQ0. congruence.*) Qed. Hint Resolve transf_instr_freshreg_trans : htlspec.

forall (fin ret st0 : reg) (l : list (node * RTL.instruction)) (s s' : st) (i : st_prop s s'), HTLMonadExtra.collectlist (transf_instr fin ret st0) l s = OK tt s' i -> st_freshreg s = st_freshreg s'

forall (fin ret st0 : reg) (l : list (node * RTL.instruction)) (s s' : st) (i : st_prop s s'), HTLMonadExtra.collectlist (transf_instr fin ret st0) l s = OK tt s' i -> st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
l:list (node * RTL.instruction)
s, s':st
i:st_prop s s'
H:HTLMonadExtra.collectlist (transf_instr fin ret0 st0) l s = OK tt s' i

st_freshreg s = st_freshreg s'
fin, ret0, st0:reg
l:list (node * RTL.instruction)
s, s':st
i:st_prop s s'
H:HTLMonadExtra.collectlist (transf_instr fin ret0 st0) l s = OK tt s' i

forall (s : state) (s' : st) (x : unit) (i : st_prop s s') (y : node * RTL.instruction), transf_instr fin ret0 st0 y s = OK x s' i -> st_freshreg s = st_freshreg s'
eauto with htlspec. Qed. Ltac rewrite_states := match goal with | [ H: ?x ?s = ?x ?s' |- _ ] => let c1 := fresh "c" in let c2 := fresh "c" in remember (?x ?s) as c1; remember (?x ?s') as c2; try subst end. Ltac inv_add_instr' H := match type of H with | ?f _ _ = OK _ _ _ => unfold f in H | ?f _ _ _ = OK _ _ _ => unfold f in H | ?f _ _ _ _ = OK _ _ _ => unfold f in H | ?f _ _ _ _ _ = OK _ _ _ => unfold f in H | ?f _ _ _ _ _ _ = OK _ _ _ => unfold f in H end; repeat unfold_match H; inversion H. Ltac inv_add_instr := match goal with | H: (if ?c then _ else _) _ = OK _ _ _ |- _ => destruct c eqn:EQN; try discriminate; inv_add_instr | H: context[add_instr_skip _ _ _] |- _ => inv_add_instr' H | H: context[add_instr_skip _ _] |- _ => monadInv H; inv_incr; inv_add_instr | H: context[add_instr _ _ _ _] |- _ => inv_add_instr' H | H: context[add_instr _ _ _] |- _ => monadInv H; inv_incr; inv_add_instr | H: context[add_branch_instr _ _ _ _ _] |- _ => inv_add_instr' H | H: context[add_branch_instr _ _ _ _] |- _ => monadInv H; inv_incr; inv_add_instr | H: context[add_node_skip _ _ _] |- _ => inv_add_instr' H | H: context[add_node_skip _ _] |- _ => monadInv H; inv_incr; inv_add_instr end. Ltac destruct_optional := match goal with H: option ?r |- _ => destruct H end.

forall (l : list (node * RTL.instruction)) (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack

forall (l : list (node * RTL.instruction)) (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
a:(node * RTL.instruction)%type
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s, s':st
i:st_prop s s'
x:unit
c:PTree.t RTL.instruction
H:(do _ <- transf_instr fin rtrn stack a; HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l) s = OK x s' i
H0:list_norepet (fst a :: map fst l)
H1:forall (pc : node) (instr : RTL.instruction), a = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:a = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr

tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s, s':st
i:st_prop s s'
x:unit
c:PTree.t RTL.instruction
H:(do _ <- match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end; HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l) s = OK x s' i
H0:list_norepet (pc1 :: map fst l)
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr

tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s, s':st
i:st_prop s s'
x:unit
c:PTree.t RTL.instruction
H:(do _ <- match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end; HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l) s = OK x s' i
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)

tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s, s':st
i:st_prop s s'
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:st
INCR:st_prop s s0
INCR0:st_prop s0 s'
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s = OK x0 s0 INCR
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s' INCR0

tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)

tr_code c pc instr (st_datapath s2) (st_controllogic s2) fin rtrn (st_st s2) stack
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
e1:pc = pc1

tr_code c pc instr (st_datapath s2) (st_controllogic s2) fin rtrn (st_st s2) stack
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
tr_code c pc instr (st_datapath s2) (st_controllogic s2) fin rtrn (st_st s2) stack
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
e1:pc = pc1

tr_code c pc instr (st_datapath s2) (st_controllogic s2) fin rtrn (st_st s2) stack
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
instr:RTL.instruction
H3:c ! pc1 = Some instr
H2:(pc1, instr1) = (pc1, instr) \/ In (pc1, instr) l
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)

tr_code c pc1 instr (st_datapath s2) (st_controllogic s2) fin rtrn (st_st s2) stack
destruct instr1 eqn:?; try discriminate; try destruct_optional; inv_add_instr; econstructor; try assumption. + destruct o with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence. + destruct o0 with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence. + inversion H2. inversion H9. rewrite H. apply tr_instr_Inop. apply Z.leb_le. assumption. eapply in_map with (f := fst) in H9. contradiction. + destruct o with pc1; destruct H16; simpl in *; rewrite AssocMap.gss in H14; eauto; congruence. + destruct o0 with pc1; destruct H16; simpl in *; rewrite AssocMap.gss in H14; eauto; congruence. + inversion H2. inversion H14. unfold nonblock. replace (st_st s4) with (st_st s2) by congruence. econstructor. apply Z.leb_le; assumption. apply EQ1. eapply in_map with (f := fst) in H14. contradiction. + destruct o with pc1; destruct H16; simpl in *; rewrite AssocMap.gss in H14; eauto; congruence. + destruct o0 with pc1; destruct H16; simpl in *; rewrite AssocMap.gss in H14; eauto; congruence. + inversion H2. inversion H14. rewrite <- e2. replace (st_st s2) with (st_st s0) by congruence. econstructor. apply Z.leb_le; assumption. apply EQ1. eapply in_map with (f := fst) in H14. contradiction. + destruct o with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence. + destruct o0 with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence. + destruct H2. * inversion H2. replace (st_st s2) with (st_st s0) by congruence. econstructor. apply Z.leb_le; assumption. eauto with htlspec. * apply in_map with (f := fst) in H2. contradiction. + destruct o with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence. + destruct o0 with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence. + destruct H2. * inversion H2. replace (st_st s2) with (st_st s0) by congruence. econstructor; try (apply Z.leb_le; apply andb_prop in EQN; apply EQN). eauto with htlspec. * apply in_map with (f := fst) in H2. contradiction. (*+ destruct o with pc1; destruct H16; simpl in *; rewrite AssocMap.gss in H14; eauto; congruence. + destruct o0 with pc1; destruct H16; simpl in *; rewrite AssocMap.gss in H14; eauto; congruence. + inversion H2. * inversion H14. constructor. congruence. * apply in_map with (f := fst) in H14. contradiction. *) + destruct o with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence. + destruct o0 with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence. + inversion H2. * inversion H9. replace (st_st s2) with (st_st s0) by congruence. eauto with htlspec. * apply in_map with (f := fst) in H9. contradiction. + destruct o with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence. + destruct o0 with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence.
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
Heqi:instr1 = RTL.Ireturn None
H1:forall (pc : node) (instr : RTL.instruction), (pc1, RTL.Ireturn None) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
instr:RTL.instruction
H3:c ! pc1 = Some instr
H2:(pc1, RTL.Ireturn None) = (pc1, instr) \/ In (pc1, instr) l
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
e1:(st_datapath s1) ! pc1 = None
Heqs:check_empty_node_datapath s1 pc1 = left e1
e2:(st_controllogic s1) ! pc1 = None
Heqs0:check_empty_node_controllogic s1 pc1 = left e2
EQ:OK tt {| st_st := st_st s1; st_freshreg := st_freshreg s1; st_freshstate := st_freshstate s1; st_scldecls := st_scldecls s1; st_arrdecls := st_arrdecls s1; st_datapath := AssocMap.set pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) (st_datapath s1); st_controllogic := AssocMap.set pc1 Vskip (st_controllogic s1) |} (add_instr_skip_state_incr s1 pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) e1 e2) = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
H10:tt = x0
H11:{| st_st := st_st s1; st_freshreg := st_freshreg s1; st_freshstate := st_freshstate s1; st_scldecls := st_scldecls s1; st_arrdecls := st_arrdecls s1; st_datapath := AssocMap.set pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) (st_datapath s1); st_controllogic := AssocMap.set pc1 Vskip (st_controllogic s1) |} = s0

tr_instr fin rtrn (st_st s2) stack instr (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) Vskip
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
Heqi:instr1 = RTL.Ireturn None
H1:forall (pc : node) (instr : RTL.instruction), (pc1, RTL.Ireturn None) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
instr:RTL.instruction
H3:c ! pc1 = Some instr
H2:(pc1, RTL.Ireturn None) = (pc1, instr) \/ In (pc1, instr) l
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
e1:(st_datapath s1) ! pc1 = None
Heqs:check_empty_node_datapath s1 pc1 = left e1
e2:(st_controllogic s1) ! pc1 = None
Heqs0:check_empty_node_controllogic s1 pc1 = left e2
EQ:OK tt {| st_st := st_st s1; st_freshreg := st_freshreg s1; st_freshstate := st_freshstate s1; st_scldecls := st_scldecls s1; st_arrdecls := st_arrdecls s1; st_datapath := AssocMap.set pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) (st_datapath s1); st_controllogic := AssocMap.set pc1 Vskip (st_controllogic s1) |} (add_instr_skip_state_incr s1 pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) e1 e2) = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
H10:tt = x0
H11:{| st_st := st_st s1; st_freshreg := st_freshreg s1; st_freshstate := st_freshstate s1; st_scldecls := st_scldecls s1; st_arrdecls := st_arrdecls s1; st_datapath := AssocMap.set pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) (st_datapath s1); st_controllogic := AssocMap.set pc1 Vskip (st_controllogic s1) |} = s0
H9:(pc1, RTL.Ireturn None) = (pc1, instr)

tr_instr fin rtrn (st_st s2) stack instr (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) Vskip
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
Heqi:instr1 = RTL.Ireturn None
H1:forall (pc : node) (instr : RTL.instruction), (pc1, RTL.Ireturn None) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
instr:RTL.instruction
H3:c ! pc1 = Some instr
H2:(pc1, RTL.Ireturn None) = (pc1, instr) \/ In (pc1, instr) l
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
e1:(st_datapath s1) ! pc1 = None
Heqs:check_empty_node_datapath s1 pc1 = left e1
e2:(st_controllogic s1) ! pc1 = None
Heqs0:check_empty_node_controllogic s1 pc1 = left e2
EQ:OK tt {| st_st := st_st s1; st_freshreg := st_freshreg s1; st_freshstate := st_freshstate s1; st_scldecls := st_scldecls s1; st_arrdecls := st_arrdecls s1; st_datapath := AssocMap.set pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) (st_datapath s1); st_controllogic := AssocMap.set pc1 Vskip (st_controllogic s1) |} (add_instr_skip_state_incr s1 pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) e1 e2) = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
H10:tt = x0
H11:{| st_st := st_st s1; st_freshreg := st_freshreg s1; st_freshstate := st_freshstate s1; st_scldecls := st_scldecls s1; st_arrdecls := st_arrdecls s1; st_datapath := AssocMap.set pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) (st_datapath s1); st_controllogic := AssocMap.set pc1 Vskip (st_controllogic s1) |} = s0
H9:In (pc1, instr) l
tr_instr fin rtrn (st_st s2) stack instr (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) Vskip
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
Heqi:instr1 = RTL.Ireturn None
H1:forall (pc : node) (instr : RTL.instruction), (pc1, RTL.Ireturn None) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
instr:RTL.instruction
H3:c ! pc1 = Some instr
H2:(pc1, RTL.Ireturn None) = (pc1, instr) \/ In (pc1, instr) l
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
e1:(st_datapath s1) ! pc1 = None
Heqs:check_empty_node_datapath s1 pc1 = left e1
e2:(st_controllogic s1) ! pc1 = None
Heqs0:check_empty_node_controllogic s1 pc1 = left e2
EQ:OK tt {| st_st := st_st s1; st_freshreg := st_freshreg s1; st_freshstate := st_freshstate s1; st_scldecls := st_scldecls s1; st_arrdecls := st_arrdecls s1; st_datapath := AssocMap.set pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) (st_datapath s1); st_controllogic := AssocMap.set pc1 Vskip (st_controllogic s1) |} (add_instr_skip_state_incr s1 pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) e1 e2) = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
H10:tt = x0
H11:{| st_st := st_st s1; st_freshreg := st_freshreg s1; st_freshstate := st_freshstate s1; st_scldecls := st_scldecls s1; st_arrdecls := st_arrdecls s1; st_datapath := AssocMap.set pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) (st_datapath s1); st_controllogic := AssocMap.set pc1 Vskip (st_controllogic s1) |} = s0
H9:(pc1, RTL.Ireturn None) = (pc1, instr)

tr_instr fin rtrn (st_st s2) stack instr (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) Vskip
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
Heqi:instr1 = RTL.Ireturn None
H1:forall (pc : node) (instr : RTL.instruction), (pc1, RTL.Ireturn None) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
instr:RTL.instruction
H3:c ! pc1 = Some instr
H2:(pc1, RTL.Ireturn None) = (pc1, instr) \/ In (pc1, instr) l
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
e1:(st_datapath s1) ! pc1 = None
Heqs:check_empty_node_datapath s1 pc1 = left e1
e2:(st_controllogic s1) ! pc1 = None
Heqs0:check_empty_node_controllogic s1 pc1 = left e2
EQ:OK tt {| st_st := st_st s1; st_freshreg := st_freshreg s1; st_freshstate := st_freshstate s1; st_scldecls := st_scldecls s1; st_arrdecls := st_arrdecls s1; st_datapath := AssocMap.set pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) (st_datapath s1); st_controllogic := AssocMap.set pc1 Vskip (st_controllogic s1) |} (add_instr_skip_state_incr s1 pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) e1 e2) = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
H10:tt = x0
H11:{| st_st := st_st s1; st_freshreg := st_freshreg s1; st_freshstate := st_freshstate s1; st_scldecls := st_scldecls s1; st_arrdecls := st_arrdecls s1; st_datapath := AssocMap.set pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) (st_datapath s1); st_controllogic := AssocMap.set pc1 Vskip (st_controllogic s1) |} = s0
H9:(pc1, RTL.Ireturn None) = (pc1, instr)
H13:RTL.Ireturn None = instr

tr_instr fin rtrn (st_st s2) stack (RTL.Ireturn None) (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) Vskip
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
Heqi:instr1 = RTL.Ireturn None
H1:forall (pc : node) (instr : RTL.instruction), (pc1, RTL.Ireturn None) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
instr:RTL.instruction
H3:c ! pc1 = Some instr
H2:(pc1, RTL.Ireturn None) = (pc1, instr) \/ In (pc1, instr) l
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
e1:(st_datapath s1) ! pc1 = None
Heqs:check_empty_node_datapath s1 pc1 = left e1
e2:(st_controllogic s1) ! pc1 = None
Heqs0:check_empty_node_controllogic s1 pc1 = left e2
EQ:OK tt {| st_st := st_st s1; st_freshreg := st_freshreg s1; st_freshstate := st_freshstate s1; st_scldecls := st_scldecls s1; st_arrdecls := st_arrdecls s1; st_datapath := AssocMap.set pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) (st_datapath s1); st_controllogic := AssocMap.set pc1 Vskip (st_controllogic s1) |} (add_instr_skip_state_incr s1 pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) e1 e2) = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
H10:tt = x0
H11:{| st_st := st_st s1; st_freshreg := st_freshreg s1; st_freshstate := st_freshstate s1; st_scldecls := st_scldecls s1; st_arrdecls := st_arrdecls s1; st_datapath := AssocMap.set pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) (st_datapath s1); st_controllogic := AssocMap.set pc1 Vskip (st_controllogic s1) |} = s0
H9:(pc1, RTL.Ireturn None) = (pc1, instr)
H13:RTL.Ireturn None = instr

tr_instr fin rtrn (st_st s0) stack (RTL.Ireturn None) (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) Vskip
eauto with htlspec.
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
Heqi:instr1 = RTL.Ireturn None
H1:forall (pc : node) (instr : RTL.instruction), (pc1, RTL.Ireturn None) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
instr:RTL.instruction
H3:c ! pc1 = Some instr
H2:(pc1, RTL.Ireturn None) = (pc1, instr) \/ In (pc1, instr) l
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
e1:(st_datapath s1) ! pc1 = None
Heqs:check_empty_node_datapath s1 pc1 = left e1
e2:(st_controllogic s1) ! pc1 = None
Heqs0:check_empty_node_controllogic s1 pc1 = left e2
EQ:OK tt {| st_st := st_st s1; st_freshreg := st_freshreg s1; st_freshstate := st_freshstate s1; st_scldecls := st_scldecls s1; st_arrdecls := st_arrdecls s1; st_datapath := AssocMap.set pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) (st_datapath s1); st_controllogic := AssocMap.set pc1 Vskip (st_controllogic s1) |} (add_instr_skip_state_incr s1 pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) e1 e2) = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
H10:tt = x0
H11:{| st_st := st_st s1; st_freshreg := st_freshreg s1; st_freshstate := st_freshstate s1; st_scldecls := st_scldecls s1; st_arrdecls := st_arrdecls s1; st_datapath := AssocMap.set pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) (st_datapath s1); st_controllogic := AssocMap.set pc1 Vskip (st_controllogic s1) |} = s0
H9:In (pc1, instr) l

tr_instr fin rtrn (st_st s2) stack instr (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) Vskip
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
Heqi:instr1 = RTL.Ireturn None
H1:forall (pc : node) (instr : RTL.instruction), (pc1, RTL.Ireturn None) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
instr:RTL.instruction
H3:c ! pc1 = Some instr
H2:(pc1, RTL.Ireturn None) = (pc1, instr) \/ In (pc1, instr) l
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
e1:(st_datapath s1) ! pc1 = None
Heqs:check_empty_node_datapath s1 pc1 = left e1
e2:(st_controllogic s1) ! pc1 = None
Heqs0:check_empty_node_controllogic s1 pc1 = left e2
EQ:OK tt {| st_st := st_st s1; st_freshreg := st_freshreg s1; st_freshstate := st_freshstate s1; st_scldecls := st_scldecls s1; st_arrdecls := st_arrdecls s1; st_datapath := AssocMap.set pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) (st_datapath s1); st_controllogic := AssocMap.set pc1 Vskip (st_controllogic s1) |} (add_instr_skip_state_incr s1 pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) e1 e2) = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
H10:tt = x0
H11:{| st_st := st_st s1; st_freshreg := st_freshreg s1; st_freshstate := st_freshstate s1; st_scldecls := st_scldecls s1; st_arrdecls := st_arrdecls s1; st_datapath := AssocMap.set pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) (st_datapath s1); st_controllogic := AssocMap.set pc1 Vskip (st_controllogic s1) |} = s0
H9:In (fst (pc1, instr)) (map fst l)

tr_instr fin rtrn (st_st s2) stack instr (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) Vskip
contradiction.
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1

tr_code c pc instr (st_datapath s2) (st_controllogic s2) fin rtrn (st_st s2) stack
eapply IHl.
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1

list_norepet (map fst l)
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
In (pc, instr) l
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1

forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
In (pc, instr) l
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr)
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1

forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
In (pc, instr) l
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr)
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
H10:pc1 = pc
H11:instr1 = instr

forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
In (pc, instr) l
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
c ! pc = Some instr
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
pc:node
instr:RTL.instruction
H2:(pc, instr) = (pc, instr)
H1:forall (pc0 : node) (instr0 : RTL.instruction), (pc, instr) = (pc0, instr0) \/ In (pc0, instr0) l -> c ! pc0 = Some instr0
H3:c ! pc = Some instr
H6:~ In pc (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc

forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
In (pc, instr) l
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1

forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
In (pc, instr) l
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
pc0:node
instr0:RTL.instruction
H9:In (pc0, instr0) l

c ! pc0 = Some instr0
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
In (pc, instr) l
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
pc0:node
instr0:RTL.instruction
H1:(pc1, instr1) = (pc0, instr0) \/ In (pc0, instr0) l -> c ! pc0 = Some instr0
pc:node
instr:RTL.instruction
H2:In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
H9:In (pc0, instr0) l

c ! pc0 = Some instr0
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
In (pc, instr) l
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
pc0:node
instr0:RTL.instruction
pc:node
instr:RTL.instruction
H2:In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
H9:In (pc0, instr0) l

(pc1, instr1) = (pc0, instr0) \/ In (pc0, instr0) l
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
pc0:node
instr0:RTL.instruction
pc:node
instr:RTL.instruction
H2:In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
H9:In (pc0, instr0) l
c ! pc0 = c ! pc0
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
In (pc, instr) l
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
pc0:node
instr0:RTL.instruction
pc:node
instr:RTL.instruction
H2:In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
H9:In (pc0, instr0) l

c ! pc0 = c ! pc0
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
In (pc, instr) l
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1

In (pc, instr) l
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr)
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1

In (pc, instr) l
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
In (pc, instr) l
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
c ! pc = Some instr
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
pc:node
instr:RTL.instruction
H1:forall (pc0 : node) (instr0 : RTL.instruction), (pc, instr) = (pc0, instr0) \/ In (pc0, instr0) l -> c ! pc0 = Some instr0
H3:c ! pc = Some instr
H6:~ In pc (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc

In (pc, instr) l
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
In (pc, instr) l
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1

In (pc, instr) l
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1
c ! pc = Some instr
pc1:node
instr1:RTL.instruction
l:list (node * RTL.instruction)
IHl:forall (fin rtrn stack : reg) (s s' : st) (i : st_prop s s') (x : unit) (c : PTree.t RTL.instruction), HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s = OK x s' i -> list_norepet (map fst l) -> (forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr) -> forall (pc : node) (instr : RTL.instruction), In (pc, instr) l -> c ! pc = Some instr -> tr_code c pc instr (st_datapath s') (st_controllogic s') fin rtrn (st_st s') stack
fin, rtrn, stack:reg
s1, s2:state
H:st_st s1 = st_st s2
H0:Ple (st_freshreg s1) (st_freshreg s2)
H4:Ple (st_freshstate s1) (st_freshstate s2)
H5:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s2) ! n = (st_datapath s1) ! n
H8:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s1) ! n
x:unit
c:PTree.t RTL.instruction
H1:forall (pc : node) (instr : RTL.instruction), (pc1, instr1) = (pc, instr) \/ In (pc, instr) l -> c ! pc = Some instr
pc:node
instr:RTL.instruction
H2:(pc1, instr1) = (pc, instr) \/ In (pc, instr) l
H3:c ! pc = Some instr
H6:~ In pc1 (map fst l)
H7:list_norepet (map fst l)
x0:unit
s0:state
e0:st_st s1 = st_st s0
p1:Ple (st_freshreg s1) (st_freshreg s0)
p2:Ple (st_freshstate s1) (st_freshstate s0)
o1:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s0) ! n = (st_datapath s1) ! n
o2:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s1) ! n
e:st_st s0 = st_st s2
p:Ple (st_freshreg s0) (st_freshreg s2)
p0:Ple (st_freshstate s0) (st_freshstate s2)
o:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
o0:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
EQ:match instr1 with | RTL.Inop n' => if Z.pos n' <=? Int.max_unsigned then add_instr pc1 n' Vskip else error (Errors.msg "State is larger than 2^32.") | RTL.Iop op args dst n' => if Z.pos n' <=? Int.max_unsigned then do instr <- translate_instr op args; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst instr) else error (Errors.msg "State is larger than 2^32.") | RTL.Iload mem addr args dst n' => if Z.pos n' <=? Int.max_unsigned then do src <- translate_arr_access mem addr args stack; do _ <- declare_reg None dst 32; add_instr pc1 n' (nonblock dst src) else error (Errors.msg "State is larger than 2^32.") | RTL.Istore mem addr args src n' => if Z.pos n' <=? Int.max_unsigned then do dst <- translate_arr_access mem addr args stack; add_instr pc1 n' (Vnonblock dst (Vvar src)) else error (Errors.msg "State is larger than 2^32.") | RTL.Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") | RTL.Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") | RTL.Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") | RTL.Icond cond args n1 n2 => if (Z.pos n1 <=? Int.max_unsigned) && (Z.pos n2 <=? Int.max_unsigned) then do e <- translate_condition cond args; add_branch_instr e pc1 n1 n2 else error (Errors.msg "State is larger than 2^32.") | RTL.Ijumptable _ _ => error (Errors.msg "Ijumptable: Case statement not supported.") | RTL.Ireturn (Some r') => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vvar r'))) | RTL.Ireturn None => add_instr_skip pc1 (Vseq (block fin (Vlit (ZToValue 1))) (block rtrn (Vlit (ZToValue 0)))) end s1 = OK x0 s0 (state_incr_intro s1 s0 e0 p1 p2 o1 o2)
EQ0:HTLMonadExtra.collectlist (transf_instr fin rtrn stack) l s0 = OK x s2 (state_incr_intro s0 s2 e p p0 o o0)
n:pc <> pc1

c ! pc = Some instr
assumption. Qed. Hint Resolve iter_expand_instr_spec : htlspec.

forall (w : option io) (x y : nat) (z : st) (a : reg) (b : nat) (c : st) (d : st_prop z c), create_arr w x y z = OK (a, b) c d -> y = b /\ a = st_freshreg z /\ st_freshreg c = Pos.succ (st_freshreg z)

forall (w : option io) (x y : nat) (z : st) (a : reg) (b : nat) (c : st) (d : st_prop z c), create_arr w x y z = OK (a, b) c d -> y = b /\ a = st_freshreg z /\ st_freshreg c = Pos.succ (st_freshreg z)
inversion 1; split; auto. Qed.

forall (a : option io) (b : nat) (s : st) (r : reg) (s' : st) (i : st_prop s s'), create_reg a b s = OK r s' i -> r = st_freshreg s /\ st_freshreg s' = Pos.succ (st_freshreg s)

forall (a : option io) (b : nat) (s : st) (r : reg) (s' : st) (i : st_prop s s'), create_reg a b s = OK r s' i -> r = st_freshreg s /\ st_freshreg s' = Pos.succ (st_freshreg s)
inversion 1; auto. Qed.

forall (f : RTL.function) (m : module), transl_module f = Errors.OK m -> tr_module f m

forall (f : RTL.function) (m : module), transl_module f = Errors.OK m -> tr_module f m
f:RTL.function
m:module

transl_module f = Errors.OK m -> tr_module f m
f:RTL.function
m:module

run_mon (max_state f) (transf_module f) = Errors.OK m -> tr_module f m
f:RTL.function
m:module

match transf_module f (max_state f) with | Error err => Errors.Error err | OK a _ _ => Errors.OK a end = Errors.OK m -> tr_module f m
f:RTL.function
m, m0:module
s':st
s:st_prop (max_state f) s'
Heqr:transf_module f (max_state f) = OK m0 s' s

Errors.OK m0 = Errors.OK m -> tr_module f m
f:RTL.function
m, m0:module
s':st
s:st_prop (max_state f) s'
Heqr:transf_module f (max_state f) = OK m0 s' s
H:Errors.OK m0 = Errors.OK m

tr_module f m
f:RTL.function
m:module
s':st
s:st_prop (max_state f) s'
Heqr:transf_module f (max_state f) = OK m s' s

tr_module f m
f:RTL.function
m:module
s':st
s:st_prop (max_state f) s'
Heqr:transf_module f (max_state f) = OK m s' s
H:st_st (max_state f) = st_st s'
H0:Ple (st_freshreg (max_state f)) (st_freshreg s')
H1:Ple (st_freshstate (max_state f)) (st_freshstate s')
H2:forall n : positive, (st_datapath (max_state f)) ! n = None \/ (st_datapath s') ! n = (st_datapath (max_state f)) ! n
H3:forall n : positive, (st_controllogic (max_state f)) ! n = None \/ (st_controllogic s') ! n = (st_controllogic (max_state f)) ! n

tr_module f m
f:RTL.function
m:module
s':st
s:st_prop (max_state f) s'
Heqr:(if stack_correct (RTL.fn_stacksize f) then do fin <- create_reg (Some Voutput) 1; do rtrn <- create_reg (Some Voutput) 32; do (stack, stack_len)<- create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)); do _ <- HTLMonadExtra.collectlist (transf_instr fin rtrn stack) (PTree.elements (RTL.fn_code f)); do _ <- HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f); do start <- create_reg (Some Vinput) 1; do rst <- create_reg (Some Vinput) 1; do clk <- create_reg (Some Vinput) 1; do current_state <- get; match zle (Z.pos (max_pc_map (st_datapath current_state))) Int.max_unsigned with | left LEDATA => match zle (Z.pos (max_pc_map (st_controllogic current_state))) Int.max_unsigned with | left LECTRL => ret {| mod_params := RTL.fn_params f; mod_datapath := st_datapath current_state; mod_controllogic := st_controllogic current_state; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st current_state; mod_stk := stack; mod_stk_len := stack_len; mod_finish := fin; mod_return := rtrn; mod_start := start; mod_reset := rst; mod_clk := clk; mod_scldecls := st_scldecls current_state; mod_arrdecls := st_arrdecls current_state; mod_wf := conj (max_pc_wf (st_controllogic current_state) LECTRL) (max_pc_wf (st_datapath current_state) LEDATA) |} | right _ => error (Errors.msg "More than 2^32 states.") end | right _ => error (Errors.msg "More than 2^32 states.") end else error (Errors.msg "Stack size misalignment.")) (max_state f) = OK m s' s
H:st_st (max_state f) = st_st s'
H0:Ple (st_freshreg (max_state f)) (st_freshreg s')
H1:Ple (st_freshstate (max_state f)) (st_freshstate s')
H2:forall n : positive, (st_datapath (max_state f)) ! n = None \/ (st_datapath s') ! n = (st_datapath (max_state f)) ! n
H3:forall n : positive, (st_controllogic (max_state f)) ! n = None \/ (st_controllogic s') ! n = (st_controllogic (max_state f)) ! n

tr_module f m
f:RTL.function
m:module
s':st
s:st_prop (max_state f) s'
Heqr:(if (0 <=? RTL.fn_stacksize f) && (RTL.fn_stacksize f <? Ptrofs.modulus) && (RTL.fn_stacksize f mod 4 =? 0) then do fin <- create_reg (Some Voutput) 1; do rtrn <- create_reg (Some Voutput) 32; do (stack, stack_len)<- create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)); do _ <- HTLMonadExtra.collectlist (transf_instr fin rtrn stack) (PTree.elements (RTL.fn_code f)); do _ <- HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f); do start <- create_reg (Some Vinput) 1; do rst <- create_reg (Some Vinput) 1; do clk <- create_reg (Some Vinput) 1; do current_state <- get; match zle (Z.pos (max_pc_map (st_datapath current_state))) Int.max_unsigned with | left LEDATA => match zle (Z.pos (max_pc_map (st_controllogic current_state))) Int.max_unsigned with | left LECTRL => ret {| mod_params := RTL.fn_params f; mod_datapath := st_datapath current_state; mod_controllogic := st_controllogic current_state; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st current_state; mod_stk := stack; mod_stk_len := stack_len; mod_finish := fin; mod_return := rtrn; mod_start := start; mod_reset := rst; mod_clk := clk; mod_scldecls := st_scldecls current_state; mod_arrdecls := st_arrdecls current_state; mod_wf := conj (max_pc_wf (st_controllogic current_state) LECTRL) (max_pc_wf (st_datapath current_state) LEDATA) |} | right _ => error (Errors.msg "More than 2^32 states.") end | right _ => error (Errors.msg "More than 2^32 states.") end else error (Errors.msg "Stack size misalignment.")) (max_state f) = OK m s' s
H:st_st (max_state f) = st_st s'
H0:Ple (st_freshreg (max_state f)) (st_freshreg s')
H1:Ple (st_freshstate (max_state f)) (st_freshstate s')
H2:forall n : positive, (st_datapath (max_state f)) ! n = None \/ (st_datapath s') ! n = (st_datapath (max_state f)) ! n
H3:forall n : positive, (st_controllogic (max_state f)) ! n = None \/ (st_controllogic s') ! n = (st_controllogic (max_state f)) ! n

tr_module f m
f:RTL.function
m:module
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
x:reg
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK x s0 INCR
x0:reg
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK x0 s1 INCR1
x1:reg
x2:nat
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (x1, x2) s2 INCR3
x3:unit
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr x x0 x1) (PTree.elements (RTL.fn_code f)) s2 = OK x3 s3 INCR5
x4:unit
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK x4 s4 INCR7
x5:reg
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK x5 s5 INCR9
x6:reg
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK x6 s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8, s8:st
INCR15:st_prop s7 s8
INCR16:st_prop s8 s'
EQ7:get s7 = OK x8 s8 INCR15
EQ9:match zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned with | left LEDATA => match zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned with | left LECTRL => ret {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := x1; mod_stk_len := x2; mod_finish := x; mod_return := x0; mod_start := x5; mod_reset := x6; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) LECTRL) (max_pc_wf (st_datapath x8) LEDATA) |} | right _ => error (Errors.msg "More than 2^32 states.") end | right _ => error (Errors.msg "More than 2^32 states.") end s8 = OK m s' INCR16

tr_module f m
f:RTL.function
m:module
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
x:reg
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK x s0 INCR
x0:reg
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK x0 s1 INCR1
x1:reg
x2:nat
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (x1, x2) s2 INCR3
x3:unit
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr x x0 x1) (PTree.elements (RTL.fn_code f)) s2 = OK x3 s3 INCR5
x4:unit
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK x4 s4 INCR7
x5:reg
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK x5 s5 INCR9
x6:reg
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK x6 s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8, s8:st
INCR15:st_prop s7 s8
INCR16:st_prop s8 s'
EQ7:get s7 = OK x8 s8 INCR15
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
EQ9:ret {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := x1; mod_stk_len := x2; mod_finish := x; mod_return := x0; mod_start := x5; mod_reset := x6; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |} s8 = OK m s' INCR16

tr_module f m
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
x:reg
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK x s0 INCR
x0:reg
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK x0 s1 INCR1
x1:reg
x2:nat
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (x1, x2) s2 INCR3
x3:unit
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr x x0 x1) (PTree.elements (RTL.fn_code f)) s2 = OK x3 s3 INCR5
x4:unit
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK x4 s4 INCR7
x5:reg
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK x5 s5 INCR9
x6:reg
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK x6 s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := x1; mod_stk_len := x2; mod_finish := x; mod_return := x0; mod_start := x5; mod_reset := x6; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
(* TODO: We should be able to fold this into the automation. *)
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
x:reg
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK x s0 INCR
x0:reg
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK x0 s1 INCR1
x1:reg
x2:nat
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (x1, x2) s2 INCR3
x3:unit
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr x x0 x1) (PTree.elements (RTL.fn_code f)) s2 = OK x3 s3 INCR5
x4:unit
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK x4 s4 INCR7
x5:reg
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK x5 s5 INCR9
x6:reg
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK x6 s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
STK_LEN:Z.to_nat (RTL.fn_stacksize f / 4) = x2 /\ x1 = st_freshreg s1 /\ st_freshreg s2 = Pos.succ (st_freshreg s1)

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := x1; mod_stk_len := x2; mod_finish := x; mod_return := x0; mod_start := x5; mod_reset := x6; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
x:reg
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK x s0 INCR
x0:reg
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK x0 s1 INCR1
x1:reg
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (x1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
x3:unit
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr x x0 x1) (PTree.elements (RTL.fn_code f)) s2 = OK x3 s3 INCR5
x4:unit
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK x4 s4 INCR7
x5:reg
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK x5 s5 INCR9
x6:reg
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK x6 s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H5:x1 = st_freshreg s1 /\ st_freshreg s2 = Pos.succ (st_freshreg s1)

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := x1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := x; mod_return := x0; mod_start := x5; mod_reset := x6; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
x:reg
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK x s0 INCR
x0:reg
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK x0 s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
x3:unit
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr x x0 (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK x3 s3 INCR5
x4:unit
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK x4 s4 INCR7
x5:reg
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK x5 s5 INCR9
x6:reg
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK x6 s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := x; mod_return := x0; mod_start := x5; mod_reset := x6; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
x:reg
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK x s0 INCR
x0:reg
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK x0 s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
x3:unit
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr x x0 (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK x3 s3 INCR5
x4:unit
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK x4 s4 INCR7
x5:reg
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK x5 s5 INCR9
x6:reg
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK x6 s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)
FIN_VAL:x = st_freshreg (max_state f) /\ st_freshreg s0 = Pos.succ (st_freshreg (max_state f))

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := x; mod_return := x0; mod_start := x5; mod_reset := x6; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
x0:reg
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK x0 s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
x3:unit
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) x0 (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK x3 s3 INCR5
x4:unit
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK x4 s4 INCR7
x5:reg
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK x5 s5 INCR9
x6:reg
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK x6 s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := x0; mod_start := x5; mod_reset := x6; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
x0:reg
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK x0 s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
x3:unit
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) x0 (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK x3 s3 INCR5
x4:unit
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK x4 s4 INCR7
x5:reg
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK x5 s5 INCR9
x6:reg
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK x6 s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))
RET_VAL:x0 = st_freshreg s0 /\ st_freshreg s1 = Pos.succ (st_freshreg s0)

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := x0; mod_start := x5; mod_reset := x6; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK (st_freshreg s0) s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
x3:unit
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) (st_freshreg s0) (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK x3 s3 INCR5
x4:unit
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK x4 s4 INCR7
x5:reg
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK x5 s5 INCR9
x6:reg
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK x6 s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))
H7:st_freshreg s1 = Pos.succ (st_freshreg s0)

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := st_freshreg s0; mod_start := x5; mod_reset := x6; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK (st_freshreg s0) s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) (st_freshreg s0) (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 INCR5
x4:unit
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK x4 s4 INCR7
x5:reg
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK x5 s5 INCR9
x6:reg
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK x6 s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))
H7:st_freshreg s1 = Pos.succ (st_freshreg s0)

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := st_freshreg s0; mod_start := x5; mod_reset := x6; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK (st_freshreg s0) s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) (st_freshreg s0) (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 INCR5
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 INCR7
x5:reg
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK x5 s5 INCR9
x6:reg
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK x6 s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))
H7:st_freshreg s1 = Pos.succ (st_freshreg s0)

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := st_freshreg s0; mod_start := x5; mod_reset := x6; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK (st_freshreg s0) s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) (st_freshreg s0) (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 INCR5
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 INCR7
x5:reg
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK x5 s5 INCR9
x6:reg
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK x6 s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))
H7:st_freshreg s1 = Pos.succ (st_freshreg s0)
TR_INSTR:st_freshreg s2 = st_freshreg s3

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := st_freshreg s0; mod_start := x5; mod_reset := x6; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK (st_freshreg s0) s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) (st_freshreg s0) (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 INCR5
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 INCR7
x5:reg
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK x5 s5 INCR9
x6:reg
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK x6 s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))
H7:st_freshreg s1 = Pos.succ (st_freshreg s0)
TR_INSTR:st_freshreg s2 = st_freshreg s3
TR_DEC:st_freshreg s3 = st_freshreg s4

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := st_freshreg s0; mod_start := x5; mod_reset := x6; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK (st_freshreg s0) s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) (st_freshreg s0) (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 INCR5
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 INCR7
x5:reg
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK x5 s5 INCR9
x6:reg
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK x6 s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))
H7:st_freshreg s1 = Pos.succ (st_freshreg s0)
TR_INSTR:st_freshreg s2 = st_freshreg s3
TR_DEC:st_freshreg s3 = st_freshreg s4
START_VAL:x5 = st_freshreg s4 /\ st_freshreg s5 = Pos.succ (st_freshreg s4)

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := st_freshreg s0; mod_start := x5; mod_reset := x6; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK (st_freshreg s0) s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) (st_freshreg s0) (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 INCR5
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 INCR7
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK (st_freshreg s4) s5 INCR9
x6:reg
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK x6 s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))
H7:st_freshreg s1 = Pos.succ (st_freshreg s0)
TR_INSTR:st_freshreg s2 = st_freshreg s3
TR_DEC:st_freshreg s3 = st_freshreg s4
H8:st_freshreg s5 = Pos.succ (st_freshreg s4)

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := st_freshreg s0; mod_start := st_freshreg s4; mod_reset := x6; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK (st_freshreg s0) s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) (st_freshreg s0) (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 INCR5
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 INCR7
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK (st_freshreg s4) s5 INCR9
x6:reg
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK x6 s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))
H7:st_freshreg s1 = Pos.succ (st_freshreg s0)
TR_INSTR:st_freshreg s2 = st_freshreg s3
TR_DEC:st_freshreg s3 = st_freshreg s4
H8:st_freshreg s5 = Pos.succ (st_freshreg s4)
RESET_VAL:x6 = st_freshreg s5 /\ st_freshreg s6 = Pos.succ (st_freshreg s5)

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := st_freshreg s0; mod_start := st_freshreg s4; mod_reset := x6; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK (st_freshreg s0) s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) (st_freshreg s0) (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 INCR5
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 INCR7
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK (st_freshreg s4) s5 INCR9
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK (st_freshreg s5) s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))
H7:st_freshreg s1 = Pos.succ (st_freshreg s0)
TR_INSTR:st_freshreg s2 = st_freshreg s3
TR_DEC:st_freshreg s3 = st_freshreg s4
H8:st_freshreg s5 = Pos.succ (st_freshreg s4)
H9:st_freshreg s6 = Pos.succ (st_freshreg s5)

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := st_freshreg s0; mod_start := st_freshreg s4; mod_reset := st_freshreg s5; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK (st_freshreg s0) s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) (st_freshreg s0) (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 INCR5
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 INCR7
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK (st_freshreg s4) s5 INCR9
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK (st_freshreg s5) s6 INCR11
x7:reg
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK x7 s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))
H7:st_freshreg s1 = Pos.succ (st_freshreg s0)
TR_INSTR:st_freshreg s2 = st_freshreg s3
TR_DEC:st_freshreg s3 = st_freshreg s4
H8:st_freshreg s5 = Pos.succ (st_freshreg s4)
H9:st_freshreg s6 = Pos.succ (st_freshreg s5)
CLK_VAL:x7 = st_freshreg s6 /\ st_freshreg s7 = Pos.succ (st_freshreg s6)

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := st_freshreg s0; mod_start := st_freshreg s4; mod_reset := st_freshreg s5; mod_clk := x7; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK (st_freshreg s0) s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) (st_freshreg s0) (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 INCR5
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 INCR7
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK (st_freshreg s4) s5 INCR9
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK (st_freshreg s5) s6 INCR11
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK (st_freshreg s6) s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))
H7:st_freshreg s1 = Pos.succ (st_freshreg s0)
TR_INSTR:st_freshreg s2 = st_freshreg s3
TR_DEC:st_freshreg s3 = st_freshreg s4
H8:st_freshreg s5 = Pos.succ (st_freshreg s4)
H9:st_freshreg s6 = Pos.succ (st_freshreg s5)
H10:st_freshreg s7 = Pos.succ (st_freshreg s6)

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := st_freshreg s0; mod_start := st_freshreg s4; mod_reset := st_freshreg s5; mod_clk := st_freshreg s6; mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK (st_freshreg s0) s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) (st_freshreg s0) (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 INCR5
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 INCR7
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK (st_freshreg s4) s5 INCR9
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK (st_freshreg s5) s6 INCR11
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK (Pos.succ (st_freshreg s5)) s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))
H7:st_freshreg s1 = Pos.succ (st_freshreg s0)
TR_INSTR:st_freshreg s2 = st_freshreg s3
TR_DEC:st_freshreg s3 = st_freshreg s4
H8:st_freshreg s5 = Pos.succ (st_freshreg s4)
H9:st_freshreg s6 = Pos.succ (st_freshreg s5)
H10:st_freshreg s7 = Pos.succ (Pos.succ (st_freshreg s5))

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := st_freshreg s0; mod_start := st_freshreg s4; mod_reset := st_freshreg s5; mod_clk := Pos.succ (st_freshreg s5); mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK (st_freshreg s0) s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) (st_freshreg s0) (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 INCR5
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 INCR7
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK (st_freshreg s4) s5 INCR9
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK (Pos.succ (st_freshreg s4)) s6 INCR11
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK (Pos.succ (Pos.succ (st_freshreg s4))) s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))
H7:st_freshreg s1 = Pos.succ (st_freshreg s0)
TR_INSTR:st_freshreg s2 = st_freshreg s3
TR_DEC:st_freshreg s3 = st_freshreg s4
H8:st_freshreg s5 = Pos.succ (st_freshreg s4)
H9:st_freshreg s6 = Pos.succ (Pos.succ (st_freshreg s4))
H10:st_freshreg s7 = Pos.succ (Pos.succ (Pos.succ (st_freshreg s4)))

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := st_freshreg s0; mod_start := st_freshreg s4; mod_reset := Pos.succ (st_freshreg s4); mod_clk := Pos.succ (Pos.succ (st_freshreg s4)); mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK (st_freshreg s0) s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) (st_freshreg s0) (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 INCR5
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 INCR7
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK (st_freshreg s2) s5 INCR9
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK (Pos.succ (st_freshreg s2)) s6 INCR11
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK (Pos.succ (Pos.succ (st_freshreg s2))) s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))
H7:st_freshreg s1 = Pos.succ (st_freshreg s0)
TR_INSTR:st_freshreg s2 = st_freshreg s3
TR_DEC:st_freshreg s3 = st_freshreg s2
H8:st_freshreg s5 = Pos.succ (st_freshreg s2)
H9:st_freshreg s6 = Pos.succ (Pos.succ (st_freshreg s2))
H10:st_freshreg s7 = Pos.succ (Pos.succ (Pos.succ (st_freshreg s2)))

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := st_freshreg s0; mod_start := st_freshreg s2; mod_reset := Pos.succ (st_freshreg s2); mod_clk := Pos.succ (Pos.succ (st_freshreg s2)); mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK (st_freshreg s0) s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (st_freshreg s1, Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) (st_freshreg s0) (st_freshreg s1)) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 INCR5
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 INCR7
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK (Pos.succ (st_freshreg s1)) s5 INCR9
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK (Pos.succ (Pos.succ (st_freshreg s1))) s6 INCR11
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK (Pos.succ (Pos.succ (Pos.succ (st_freshreg s1)))) s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (st_freshreg s1)
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))
H7:st_freshreg s1 = Pos.succ (st_freshreg s0)
TR_INSTR:Pos.succ (st_freshreg s1) = st_freshreg s3
TR_DEC:st_freshreg s3 = Pos.succ (st_freshreg s1)
H8:st_freshreg s5 = Pos.succ (Pos.succ (st_freshreg s1))
H9:st_freshreg s6 = Pos.succ (Pos.succ (Pos.succ (st_freshreg s1)))
H10:st_freshreg s7 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (st_freshreg s1))))

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := st_freshreg s1; mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := st_freshreg s0; mod_start := Pos.succ (st_freshreg s1); mod_reset := Pos.succ (Pos.succ (st_freshreg s1)); mod_clk := Pos.succ (Pos.succ (Pos.succ (st_freshreg s1))); mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK (st_freshreg s0) s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (Pos.succ (st_freshreg s0), Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) (st_freshreg s0) (Pos.succ (st_freshreg s0))) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 INCR5
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 INCR7
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK (Pos.succ (Pos.succ (st_freshreg s0))) s5 INCR9
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK (Pos.succ (Pos.succ (Pos.succ (st_freshreg s0)))) s6 INCR11
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK (Pos.succ (Pos.succ (Pos.succ (Pos.succ (st_freshreg s0))))) s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (Pos.succ (st_freshreg s0))
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))
H7:st_freshreg s1 = Pos.succ (st_freshreg s0)
TR_INSTR:Pos.succ (Pos.succ (st_freshreg s0)) = st_freshreg s3
TR_DEC:st_freshreg s3 = Pos.succ (Pos.succ (st_freshreg s0))
H8:st_freshreg s5 = Pos.succ (Pos.succ (Pos.succ (st_freshreg s0)))
H9:st_freshreg s6 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (st_freshreg s0))))
H10:st_freshreg s7 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (st_freshreg s0)))))

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := Pos.succ (st_freshreg s0); mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := st_freshreg s0; mod_start := Pos.succ (Pos.succ (st_freshreg s0)); mod_reset := Pos.succ (Pos.succ (Pos.succ (st_freshreg s0))); mod_clk := Pos.succ (Pos.succ (Pos.succ (Pos.succ (st_freshreg s0)))); mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (st_freshreg (max_state f)) s0 INCR
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK (Pos.succ (st_freshreg (max_state f))) s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (Pos.succ (Pos.succ (st_freshreg (max_state f))), Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (st_freshreg (max_state f)) (Pos.succ (st_freshreg (max_state f))) (Pos.succ (Pos.succ (st_freshreg (max_state f))))) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 INCR5
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 INCR7
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK (Pos.succ (Pos.succ (Pos.succ (st_freshreg (max_state f))))) s5 INCR9
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK (Pos.succ (Pos.succ (Pos.succ (Pos.succ (st_freshreg (max_state f)))))) s6 INCR11
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (st_freshreg (max_state f))))))) s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (Pos.succ (Pos.succ (st_freshreg (max_state f))))
H5:st_freshreg s0 = Pos.succ (st_freshreg (max_state f))
H7:st_freshreg s1 = Pos.succ (Pos.succ (st_freshreg (max_state f)))
TR_INSTR:Pos.succ (Pos.succ (Pos.succ (st_freshreg (max_state f)))) = st_freshreg s3
TR_DEC:st_freshreg s3 = Pos.succ (Pos.succ (Pos.succ (st_freshreg (max_state f))))
H8:st_freshreg s5 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (st_freshreg (max_state f)))))
H9:st_freshreg s6 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (st_freshreg (max_state f))))))
H10:st_freshreg s7 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (st_freshreg (max_state f)))))))

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := Pos.succ (Pos.succ (st_freshreg (max_state f))); mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := st_freshreg (max_state f); mod_return := Pos.succ (st_freshreg (max_state f)); mod_start := Pos.succ (Pos.succ (Pos.succ (st_freshreg (max_state f)))); mod_reset := Pos.succ (Pos.succ (Pos.succ (Pos.succ (st_freshreg (max_state f))))); mod_clk := Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (st_freshreg (max_state f)))))); mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s':st
s:st_prop (max_state f) s'
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s'
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s')
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s')
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s') ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s') ! n = (AssocMap.empty stmnt) ! n
s0:st
INCR:st_prop (max_state f) s0
INCR0:st_prop s0 s'
EQ:create_reg (Some Voutput) 1 (max_state f) = OK (Pos.succ (Pos.succ (RTL.max_reg_function f))) s0 INCR
s1:st
INCR1:st_prop s0 s1
INCR2:st_prop s1 s'
EQ1:create_reg (Some Voutput) 32 s0 = OK (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) s1 INCR1
s2:st
INCR3:st_prop s1 s2
INCR4:st_prop s2 s'
EQ0:create_arr None 32 (Z.to_nat (RTL.fn_stacksize f / 4)) s1 = OK (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))), Z.to_nat (RTL.fn_stacksize f / 4)) s2 INCR3
s3:st
INCR5:st_prop s2 s3
INCR6:st_prop s3 s'
EQ2:HTLMonadExtra.collectlist (transf_instr (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 INCR5
s4:st
INCR7:st_prop s3 s4
INCR8:st_prop s4 s'
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 INCR7
s5:st
INCR9:st_prop s4 s5
INCR10:st_prop s5 s'
EQ4:create_reg (Some Vinput) 1 s4 = OK (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))) s5 INCR9
s6:st
INCR11:st_prop s5 s6
INCR12:st_prop s6 s'
EQ5:create_reg (Some Vinput) 1 s5 = OK (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))) s6 INCR11
s7:st
INCR13:st_prop s6 s7
INCR14:st_prop s7 s'
EQ6:create_reg (Some Vinput) 1 s6 = OK (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))))) s7 INCR13
x8:st
INCR15:st_prop s7 s'
EQ7:get s7 = OK x8 s' INCR15
INCR16:st_prop s' s'
l:Z.pos (max_pc_map (st_datapath x8)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath x8))) Int.max_unsigned = left l
l0:Z.pos (max_pc_map (st_controllogic x8)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic x8))) Int.max_unsigned = left l0
H6:st_freshreg s2 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H5:st_freshreg s0 = Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))
H7:st_freshreg s1 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))
TR_INSTR:Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))) = st_freshreg s3
TR_DEC:st_freshreg s3 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H8:st_freshreg s5 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))
H9:st_freshreg s6 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))))
H10:st_freshreg s7 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))))

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath x8; mod_controllogic := st_controllogic x8; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st x8; mod_stk := Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))); mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := Pos.succ (Pos.succ (RTL.max_reg_function f)); mod_return := Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))); mod_start := Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))); mod_reset := Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))); mod_clk := Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))); mod_scldecls := st_scldecls x8; mod_arrdecls := st_arrdecls x8; mod_wf := conj (max_pc_wf (st_controllogic x8) l0) (max_pc_wf (st_datapath x8) l) |}
f:RTL.function
s1, s10:state
H96:st_st s1 = st_st s10
H97:Ple (st_freshreg s1) (st_freshreg s10)
H98:Ple (st_freshstate s1) (st_freshstate s10)
H99:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s10) ! n = (st_datapath s1) ! n
H100:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s1) ! n
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s10
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s10)
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s10)
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s10) ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s10) ! n = (AssocMap.empty stmnt) ! n
s11:state
H91:st_st s1 = st_st s11
H92:Ple (st_freshreg s1) (st_freshreg s11)
H93:Ple (st_freshstate s1) (st_freshstate s11)
H94:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s11) ! n = (st_datapath s1) ! n
H95:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s11) ! n = (st_controllogic s1) ! n
H86:st_st s11 = st_st s10
H87:Ple (st_freshreg s11) (st_freshreg s10)
H88:Ple (st_freshstate s11) (st_freshstate s10)
H89:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s10) ! n = (st_datapath s11) ! n
H90:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s11) ! n
EQ:st_datapath s1 = st_datapath s11
s0:state
H81:st_st s11 = st_st s0
H82:Ple (st_freshreg s11) (st_freshreg s0)
H83:Ple (st_freshstate s11) (st_freshstate s0)
H84:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s0) ! n = (st_datapath s11) ! n
H85:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s11) ! n
H76:st_st s0 = st_st s10
H77:Ple (st_freshreg s0) (st_freshreg s10)
H78:Ple (st_freshstate s0) (st_freshstate s10)
H79:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s10) ! n = (st_datapath s0) ! n
H80:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s0) ! n
EQ1:st_datapath s11 = st_datapath s0
s2:state
H71:st_st s0 = st_st s2
H72:Ple (st_freshreg s0) (st_freshreg s2)
H73:Ple (st_freshstate s0) (st_freshstate s2)
H74:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
H75:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
H66:st_st s2 = st_st s10
H67:Ple (st_freshreg s2) (st_freshreg s10)
H68:Ple (st_freshstate s2) (st_freshstate s10)
H69:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s10) ! n = (st_datapath s2) ! n
H70:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s2) ! n
EQ0:st_datapath s0 = st_datapath s2
s3:state
e0:st_st s2 = st_st s3
p1:Ple (st_freshreg s2) (st_freshreg s3)
p2:Ple (st_freshstate s2) (st_freshstate s3)
o1:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s3) ! n = (st_datapath s2) ! n
o2:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s3) ! n = (st_controllogic s2) ! n
H61:st_st s3 = st_st s10
H62:Ple (st_freshreg s3) (st_freshreg s10)
H63:Ple (st_freshstate s3) (st_freshstate s10)
H64:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s10) ! n = (st_datapath s3) ! n
H65:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s3) ! n
EQ2:HTLMonadExtra.collectlist (transf_instr (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 (state_incr_intro s2 s3 e0 p1 p2 o1 o2)
s4:state
e:st_st s3 = st_st s4
p:Ple (st_freshreg s3) (st_freshreg s4)
p0:Ple (st_freshstate s3) (st_freshstate s4)
o:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s4) ! n = (st_datapath s3) ! n
o0:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s4) ! n = (st_controllogic s3) ! n
H56:st_st s4 = st_st s10
H57:Ple (st_freshreg s4) (st_freshreg s10)
H58:Ple (st_freshstate s4) (st_freshstate s10)
H59:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s10) ! n = (st_datapath s4) ! n
H60:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s4) ! n
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 (state_incr_intro s3 s4 e p p0 o o0)
s5:state
H51:st_st s4 = st_st s5
H52:Ple (st_freshreg s4) (st_freshreg s5)
H53:Ple (st_freshstate s4) (st_freshstate s5)
H54:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s5) ! n = (st_datapath s4) ! n
H55:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s5) ! n = (st_controllogic s4) ! n
H46:st_st s5 = st_st s10
H47:Ple (st_freshreg s5) (st_freshreg s10)
H48:Ple (st_freshstate s5) (st_freshstate s10)
H49:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s10) ! n = (st_datapath s5) ! n
H50:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s5) ! n
EQ4:st_datapath s4 = st_datapath s5
s6:state
H41:st_st s5 = st_st s6
H42:Ple (st_freshreg s5) (st_freshreg s6)
H43:Ple (st_freshstate s5) (st_freshstate s6)
H44:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s6) ! n = (st_datapath s5) ! n
H45:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s6) ! n = (st_controllogic s5) ! n
H36:st_st s6 = st_st s10
H37:Ple (st_freshreg s6) (st_freshreg s10)
H38:Ple (st_freshstate s6) (st_freshstate s10)
H39:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H40:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
EQ5:st_datapath s5 = st_datapath s6
H31:st_st s6 = st_st s10
H32:Ple (st_freshreg s6) (st_freshreg s10)
H33:Ple (st_freshstate s6) (st_freshstate s10)
H34:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H35:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
s9:state
H26:st_st s9 = st_st s10
H27:Ple (st_freshreg s9) (st_freshreg s10)
H28:Ple (st_freshstate s9) (st_freshstate s10)
H29:forall n : positive, (st_datapath s9) ! n = None \/ (st_datapath s10) ! n = (st_datapath s9) ! n
H30:forall n : positive, (st_controllogic s9) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s9) ! n
EQ6:st_datapath s6 = st_datapath s10
s8:state
H21:st_st s8 = st_st s10
H22:Ple (st_freshreg s8) (st_freshreg s10)
H23:Ple (st_freshstate s8) (st_freshstate s10)
H24:forall n : positive, (st_datapath s8) ! n = None \/ (st_datapath s10) ! n = (st_datapath s8) ! n
H25:forall n : positive, (st_controllogic s8) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s8) ! n
s7:state
H16:st_st s7 = st_st s10
H17:Ple (st_freshreg s7) (st_freshreg s10)
H18:Ple (st_freshstate s7) (st_freshstate s10)
H19:forall n : positive, (st_datapath s7) ! n = None \/ (st_datapath s10) ! n = (st_datapath s7) ! n
H20:forall n : positive, (st_controllogic s7) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s7) ! n
l0:Z.pos (max_pc_map (st_controllogic s10)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic s10))) Int.max_unsigned = left l0
l:Z.pos (max_pc_map (st_datapath s10)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath s10))) Int.max_unsigned = left l
H6:st_freshreg s2 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H5:st_freshreg s11 = Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))
H7:st_freshreg s0 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))
TR_INSTR:Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))) = st_freshreg s3
TR_DEC:st_freshreg s3 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H8:st_freshreg s5 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))
H9:st_freshreg s6 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))))
H10:st_freshreg s10 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))))
H4:st_controllogic s6 = st_controllogic s10
H11:st_controllogic s5 = st_controllogic s6
H12:st_controllogic s4 = st_controllogic s5
H13:st_controllogic s11 = st_controllogic s0
H14:st_controllogic s1 = st_controllogic s11
H15:st_controllogic s0 = st_controllogic s2

tr_module f {| mod_params := RTL.fn_params f; mod_datapath := st_datapath s10; mod_controllogic := st_controllogic s10; mod_entrypoint := RTL.fn_entrypoint f; mod_st := st_st s10; mod_stk := Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))); mod_stk_len := Z.to_nat (RTL.fn_stacksize f / 4); mod_finish := Pos.succ (Pos.succ (RTL.max_reg_function f)); mod_return := Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))); mod_start := Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))); mod_reset := Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))); mod_clk := Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))); mod_scldecls := st_scldecls s10; mod_arrdecls := st_arrdecls s10; mod_wf := conj (max_pc_wf (st_controllogic s10) l0) (max_pc_wf (st_datapath s10) l) |}
f:RTL.function
s1, s10:state
H96:st_st s1 = st_st s10
H97:Ple (st_freshreg s1) (st_freshreg s10)
H98:Ple (st_freshstate s1) (st_freshstate s10)
H99:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s10) ! n = (st_datapath s1) ! n
H100:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s1) ! n
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s10
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s10)
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s10)
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s10) ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s10) ! n = (AssocMap.empty stmnt) ! n
s11:state
H91:st_st s1 = st_st s11
H92:Ple (st_freshreg s1) (st_freshreg s11)
H93:Ple (st_freshstate s1) (st_freshstate s11)
H94:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s11) ! n = (st_datapath s1) ! n
H95:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s11) ! n = (st_controllogic s1) ! n
H86:st_st s11 = st_st s10
H87:Ple (st_freshreg s11) (st_freshreg s10)
H88:Ple (st_freshstate s11) (st_freshstate s10)
H89:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s10) ! n = (st_datapath s11) ! n
H90:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s11) ! n
EQ:st_datapath s1 = st_datapath s11
s0:state
H81:st_st s11 = st_st s0
H82:Ple (st_freshreg s11) (st_freshreg s0)
H83:Ple (st_freshstate s11) (st_freshstate s0)
H84:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s0) ! n = (st_datapath s11) ! n
H85:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s11) ! n
H76:st_st s0 = st_st s10
H77:Ple (st_freshreg s0) (st_freshreg s10)
H78:Ple (st_freshstate s0) (st_freshstate s10)
H79:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s10) ! n = (st_datapath s0) ! n
H80:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s0) ! n
EQ1:st_datapath s11 = st_datapath s0
s2:state
H71:st_st s0 = st_st s2
H72:Ple (st_freshreg s0) (st_freshreg s2)
H73:Ple (st_freshstate s0) (st_freshstate s2)
H74:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
H75:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
H66:st_st s2 = st_st s10
H67:Ple (st_freshreg s2) (st_freshreg s10)
H68:Ple (st_freshstate s2) (st_freshstate s10)
H69:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s10) ! n = (st_datapath s2) ! n
H70:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s2) ! n
EQ0:st_datapath s0 = st_datapath s2
s3:state
e0:st_st s2 = st_st s3
p1:Ple (st_freshreg s2) (st_freshreg s3)
p2:Ple (st_freshstate s2) (st_freshstate s3)
o1:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s3) ! n = (st_datapath s2) ! n
o2:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s3) ! n = (st_controllogic s2) ! n
H61:st_st s3 = st_st s10
H62:Ple (st_freshreg s3) (st_freshreg s10)
H63:Ple (st_freshstate s3) (st_freshstate s10)
H64:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s10) ! n = (st_datapath s3) ! n
H65:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s3) ! n
EQ2:HTLMonadExtra.collectlist (transf_instr (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 (state_incr_intro s2 s3 e0 p1 p2 o1 o2)
s4:state
e:st_st s3 = st_st s4
p:Ple (st_freshreg s3) (st_freshreg s4)
p0:Ple (st_freshstate s3) (st_freshstate s4)
o:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s4) ! n = (st_datapath s3) ! n
o0:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s4) ! n = (st_controllogic s3) ! n
H56:st_st s4 = st_st s10
H57:Ple (st_freshreg s4) (st_freshreg s10)
H58:Ple (st_freshstate s4) (st_freshstate s10)
H59:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s10) ! n = (st_datapath s4) ! n
H60:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s4) ! n
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 (state_incr_intro s3 s4 e p p0 o o0)
s5:state
H51:st_st s4 = st_st s5
H52:Ple (st_freshreg s4) (st_freshreg s5)
H53:Ple (st_freshstate s4) (st_freshstate s5)
H54:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s5) ! n = (st_datapath s4) ! n
H55:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s5) ! n = (st_controllogic s4) ! n
H46:st_st s5 = st_st s10
H47:Ple (st_freshreg s5) (st_freshreg s10)
H48:Ple (st_freshstate s5) (st_freshstate s10)
H49:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s10) ! n = (st_datapath s5) ! n
H50:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s5) ! n
EQ4:st_datapath s4 = st_datapath s5
s6:state
H41:st_st s5 = st_st s6
H42:Ple (st_freshreg s5) (st_freshreg s6)
H43:Ple (st_freshstate s5) (st_freshstate s6)
H44:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s6) ! n = (st_datapath s5) ! n
H45:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s6) ! n = (st_controllogic s5) ! n
H36:st_st s6 = st_st s10
H37:Ple (st_freshreg s6) (st_freshreg s10)
H38:Ple (st_freshstate s6) (st_freshstate s10)
H39:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H40:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
EQ5:st_datapath s5 = st_datapath s6
H31:st_st s6 = st_st s10
H32:Ple (st_freshreg s6) (st_freshreg s10)
H33:Ple (st_freshstate s6) (st_freshstate s10)
H34:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H35:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
s9:state
H26:st_st s9 = st_st s10
H27:Ple (st_freshreg s9) (st_freshreg s10)
H28:Ple (st_freshstate s9) (st_freshstate s10)
H29:forall n : positive, (st_datapath s9) ! n = None \/ (st_datapath s10) ! n = (st_datapath s9) ! n
H30:forall n : positive, (st_controllogic s9) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s9) ! n
EQ6:st_datapath s6 = st_datapath s10
s8:state
H21:st_st s8 = st_st s10
H22:Ple (st_freshreg s8) (st_freshreg s10)
H23:Ple (st_freshstate s8) (st_freshstate s10)
H24:forall n : positive, (st_datapath s8) ! n = None \/ (st_datapath s10) ! n = (st_datapath s8) ! n
H25:forall n : positive, (st_controllogic s8) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s8) ! n
s7:state
H16:st_st s7 = st_st s10
H17:Ple (st_freshreg s7) (st_freshreg s10)
H18:Ple (st_freshstate s7) (st_freshstate s10)
H19:forall n : positive, (st_datapath s7) ! n = None \/ (st_datapath s10) ! n = (st_datapath s7) ! n
H20:forall n : positive, (st_controllogic s7) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s7) ! n
l0:Z.pos (max_pc_map (st_controllogic s10)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic s10))) Int.max_unsigned = left l0
l:Z.pos (max_pc_map (st_datapath s10)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath s10))) Int.max_unsigned = left l
H6:st_freshreg s2 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H5:st_freshreg s11 = Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))
H7:st_freshreg s0 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))
TR_INSTR:Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))) = st_freshreg s3
TR_DEC:st_freshreg s3 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H8:st_freshreg s5 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))
H9:st_freshreg s6 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))))
H10:st_freshreg s10 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))))
H4:st_controllogic s6 = st_controllogic s10
H11:st_controllogic s5 = st_controllogic s6
H12:st_controllogic s4 = st_controllogic s5
H13:st_controllogic s11 = st_controllogic s0
H14:st_controllogic s1 = st_controllogic s11
H15:st_controllogic s0 = st_controllogic s2

forall (pc : positive) (i : RTL.instruction), (RTL.fn_code f) ! pc = Some i -> tr_code (RTL.fn_code f) pc i (st_datapath s10) (st_controllogic s10) (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (st_st s10) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
f:RTL.function
s1, s10:state
H96:st_st s1 = st_st s10
H97:Ple (st_freshreg s1) (st_freshreg s10)
H98:Ple (st_freshstate s1) (st_freshstate s10)
H99:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s10) ! n = (st_datapath s1) ! n
H100:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s1) ! n
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s10
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s10)
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s10)
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s10) ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s10) ! n = (AssocMap.empty stmnt) ! n
s11:state
H91:st_st s1 = st_st s11
H92:Ple (st_freshreg s1) (st_freshreg s11)
H93:Ple (st_freshstate s1) (st_freshstate s11)
H94:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s11) ! n = (st_datapath s1) ! n
H95:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s11) ! n = (st_controllogic s1) ! n
H86:st_st s11 = st_st s10
H87:Ple (st_freshreg s11) (st_freshreg s10)
H88:Ple (st_freshstate s11) (st_freshstate s10)
H89:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s10) ! n = (st_datapath s11) ! n
H90:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s11) ! n
EQ:st_datapath s1 = st_datapath s11
s0:state
H81:st_st s11 = st_st s0
H82:Ple (st_freshreg s11) (st_freshreg s0)
H83:Ple (st_freshstate s11) (st_freshstate s0)
H84:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s0) ! n = (st_datapath s11) ! n
H85:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s11) ! n
H76:st_st s0 = st_st s10
H77:Ple (st_freshreg s0) (st_freshreg s10)
H78:Ple (st_freshstate s0) (st_freshstate s10)
H79:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s10) ! n = (st_datapath s0) ! n
H80:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s0) ! n
EQ1:st_datapath s11 = st_datapath s0
s2:state
H71:st_st s0 = st_st s2
H72:Ple (st_freshreg s0) (st_freshreg s2)
H73:Ple (st_freshstate s0) (st_freshstate s2)
H74:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
H75:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
H66:st_st s2 = st_st s10
H67:Ple (st_freshreg s2) (st_freshreg s10)
H68:Ple (st_freshstate s2) (st_freshstate s10)
H69:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s10) ! n = (st_datapath s2) ! n
H70:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s2) ! n
EQ0:st_datapath s0 = st_datapath s2
s3:state
e0:st_st s2 = st_st s3
p1:Ple (st_freshreg s2) (st_freshreg s3)
p2:Ple (st_freshstate s2) (st_freshstate s3)
o1:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s3) ! n = (st_datapath s2) ! n
o2:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s3) ! n = (st_controllogic s2) ! n
H61:st_st s3 = st_st s10
H62:Ple (st_freshreg s3) (st_freshreg s10)
H63:Ple (st_freshstate s3) (st_freshstate s10)
H64:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s10) ! n = (st_datapath s3) ! n
H65:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s3) ! n
EQ2:HTLMonadExtra.collectlist (transf_instr (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 (state_incr_intro s2 s3 e0 p1 p2 o1 o2)
s4:state
e:st_st s3 = st_st s4
p:Ple (st_freshreg s3) (st_freshreg s4)
p0:Ple (st_freshstate s3) (st_freshstate s4)
o:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s4) ! n = (st_datapath s3) ! n
o0:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s4) ! n = (st_controllogic s3) ! n
H56:st_st s4 = st_st s10
H57:Ple (st_freshreg s4) (st_freshreg s10)
H58:Ple (st_freshstate s4) (st_freshstate s10)
H59:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s10) ! n = (st_datapath s4) ! n
H60:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s4) ! n
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 (state_incr_intro s3 s4 e p p0 o o0)
s5:state
H51:st_st s4 = st_st s5
H52:Ple (st_freshreg s4) (st_freshreg s5)
H53:Ple (st_freshstate s4) (st_freshstate s5)
H54:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s5) ! n = (st_datapath s4) ! n
H55:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s5) ! n = (st_controllogic s4) ! n
H46:st_st s5 = st_st s10
H47:Ple (st_freshreg s5) (st_freshreg s10)
H48:Ple (st_freshstate s5) (st_freshstate s10)
H49:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s10) ! n = (st_datapath s5) ! n
H50:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s5) ! n
EQ4:st_datapath s4 = st_datapath s5
s6:state
H41:st_st s5 = st_st s6
H42:Ple (st_freshreg s5) (st_freshreg s6)
H43:Ple (st_freshstate s5) (st_freshstate s6)
H44:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s6) ! n = (st_datapath s5) ! n
H45:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s6) ! n = (st_controllogic s5) ! n
H36:st_st s6 = st_st s10
H37:Ple (st_freshreg s6) (st_freshreg s10)
H38:Ple (st_freshstate s6) (st_freshstate s10)
H39:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H40:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
EQ5:st_datapath s5 = st_datapath s6
H31:st_st s6 = st_st s10
H32:Ple (st_freshreg s6) (st_freshreg s10)
H33:Ple (st_freshstate s6) (st_freshstate s10)
H34:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H35:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
s9:state
H26:st_st s9 = st_st s10
H27:Ple (st_freshreg s9) (st_freshreg s10)
H28:Ple (st_freshstate s9) (st_freshstate s10)
H29:forall n : positive, (st_datapath s9) ! n = None \/ (st_datapath s10) ! n = (st_datapath s9) ! n
H30:forall n : positive, (st_controllogic s9) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s9) ! n
EQ6:st_datapath s6 = st_datapath s10
s8:state
H21:st_st s8 = st_st s10
H22:Ple (st_freshreg s8) (st_freshreg s10)
H23:Ple (st_freshstate s8) (st_freshstate s10)
H24:forall n : positive, (st_datapath s8) ! n = None \/ (st_datapath s10) ! n = (st_datapath s8) ! n
H25:forall n : positive, (st_controllogic s8) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s8) ! n
s7:state
H16:st_st s7 = st_st s10
H17:Ple (st_freshreg s7) (st_freshreg s10)
H18:Ple (st_freshstate s7) (st_freshstate s10)
H19:forall n : positive, (st_datapath s7) ! n = None \/ (st_datapath s10) ! n = (st_datapath s7) ! n
H20:forall n : positive, (st_controllogic s7) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s7) ! n
l0:Z.pos (max_pc_map (st_controllogic s10)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic s10))) Int.max_unsigned = left l0
l:Z.pos (max_pc_map (st_datapath s10)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath s10))) Int.max_unsigned = left l
H6:st_freshreg s2 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H5:st_freshreg s11 = Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))
H7:st_freshreg s0 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))
TR_INSTR:Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))) = st_freshreg s3
TR_DEC:st_freshreg s3 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H8:st_freshreg s5 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))
H9:st_freshreg s6 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))))
H10:st_freshreg s10 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))))
H4:st_controllogic s6 = st_controllogic s10
H11:st_controllogic s5 = st_controllogic s6
H12:st_controllogic s4 = st_controllogic s5
H13:st_controllogic s11 = st_controllogic s0
H14:st_controllogic s1 = st_controllogic s11
H15:st_controllogic s0 = st_controllogic s2
pc:positive
i:RTL.instruction
H101:(RTL.fn_code f) ! pc = Some i

tr_code (RTL.fn_code f) pc i (st_datapath s10) (st_controllogic s10) (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (st_st s10) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
f:RTL.function
s1, s10:state
H96:st_st s1 = st_st s10
H97:Ple (st_freshreg s1) (st_freshreg s10)
H98:Ple (st_freshstate s1) (st_freshstate s10)
H99:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s10) ! n = (st_datapath s1) ! n
H100:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s1) ! n
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s10
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s10)
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s10)
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s10) ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s10) ! n = (AssocMap.empty stmnt) ! n
s11:state
H91:st_st s1 = st_st s11
H92:Ple (st_freshreg s1) (st_freshreg s11)
H93:Ple (st_freshstate s1) (st_freshstate s11)
H94:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s11) ! n = (st_datapath s1) ! n
H95:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s11) ! n = (st_controllogic s1) ! n
H86:st_st s11 = st_st s10
H87:Ple (st_freshreg s11) (st_freshreg s10)
H88:Ple (st_freshstate s11) (st_freshstate s10)
H89:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s10) ! n = (st_datapath s11) ! n
H90:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s11) ! n
EQ:st_datapath s1 = st_datapath s11
s0:state
H81:st_st s11 = st_st s0
H82:Ple (st_freshreg s11) (st_freshreg s0)
H83:Ple (st_freshstate s11) (st_freshstate s0)
H84:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s0) ! n = (st_datapath s11) ! n
H85:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s11) ! n
H76:st_st s0 = st_st s10
H77:Ple (st_freshreg s0) (st_freshreg s10)
H78:Ple (st_freshstate s0) (st_freshstate s10)
H79:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s10) ! n = (st_datapath s0) ! n
H80:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s0) ! n
EQ1:st_datapath s11 = st_datapath s0
s2:state
H71:st_st s0 = st_st s2
H72:Ple (st_freshreg s0) (st_freshreg s2)
H73:Ple (st_freshstate s0) (st_freshstate s2)
H74:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
H75:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
H66:st_st s2 = st_st s10
H67:Ple (st_freshreg s2) (st_freshreg s10)
H68:Ple (st_freshstate s2) (st_freshstate s10)
H69:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s10) ! n = (st_datapath s2) ! n
H70:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s2) ! n
EQ0:st_datapath s0 = st_datapath s2
s3:state
e0:st_st s2 = st_st s3
p1:Ple (st_freshreg s2) (st_freshreg s3)
p2:Ple (st_freshstate s2) (st_freshstate s3)
o1:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s3) ! n = (st_datapath s2) ! n
o2:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s3) ! n = (st_controllogic s2) ! n
H61:st_st s3 = st_st s10
H62:Ple (st_freshreg s3) (st_freshreg s10)
H63:Ple (st_freshstate s3) (st_freshstate s10)
H64:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s10) ! n = (st_datapath s3) ! n
H65:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s3) ! n
EQ2:HTLMonadExtra.collectlist (transf_instr (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 (state_incr_intro s2 s3 e0 p1 p2 o1 o2)
s4:state
e:st_st s3 = st_st s4
p:Ple (st_freshreg s3) (st_freshreg s4)
p0:Ple (st_freshstate s3) (st_freshstate s4)
o:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s4) ! n = (st_datapath s3) ! n
o0:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s4) ! n = (st_controllogic s3) ! n
H56:st_st s4 = st_st s10
H57:Ple (st_freshreg s4) (st_freshreg s10)
H58:Ple (st_freshstate s4) (st_freshstate s10)
H59:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s10) ! n = (st_datapath s4) ! n
H60:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s4) ! n
EQ3:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 (state_incr_intro s3 s4 e p p0 o o0)
s5:state
H51:st_st s4 = st_st s5
H52:Ple (st_freshreg s4) (st_freshreg s5)
H53:Ple (st_freshstate s4) (st_freshstate s5)
H54:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s5) ! n = (st_datapath s4) ! n
H55:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s5) ! n = (st_controllogic s4) ! n
H46:st_st s5 = st_st s10
H47:Ple (st_freshreg s5) (st_freshreg s10)
H48:Ple (st_freshstate s5) (st_freshstate s10)
H49:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s10) ! n = (st_datapath s5) ! n
H50:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s5) ! n
EQ4:st_datapath s4 = st_datapath s5
s6:state
H41:st_st s5 = st_st s6
H42:Ple (st_freshreg s5) (st_freshreg s6)
H43:Ple (st_freshstate s5) (st_freshstate s6)
H44:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s6) ! n = (st_datapath s5) ! n
H45:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s6) ! n = (st_controllogic s5) ! n
H36:st_st s6 = st_st s10
H37:Ple (st_freshreg s6) (st_freshreg s10)
H38:Ple (st_freshstate s6) (st_freshstate s10)
H39:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H40:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
EQ5:st_datapath s5 = st_datapath s6
H31:st_st s6 = st_st s10
H32:Ple (st_freshreg s6) (st_freshreg s10)
H33:Ple (st_freshstate s6) (st_freshstate s10)
H34:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H35:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
s9:state
H26:st_st s9 = st_st s10
H27:Ple (st_freshreg s9) (st_freshreg s10)
H28:Ple (st_freshstate s9) (st_freshstate s10)
H29:forall n : positive, (st_datapath s9) ! n = None \/ (st_datapath s10) ! n = (st_datapath s9) ! n
H30:forall n : positive, (st_controllogic s9) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s9) ! n
EQ6:st_datapath s6 = st_datapath s10
s8:state
H21:st_st s8 = st_st s10
H22:Ple (st_freshreg s8) (st_freshreg s10)
H23:Ple (st_freshstate s8) (st_freshstate s10)
H24:forall n : positive, (st_datapath s8) ! n = None \/ (st_datapath s10) ! n = (st_datapath s8) ! n
H25:forall n : positive, (st_controllogic s8) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s8) ! n
s7:state
H16:st_st s7 = st_st s10
H17:Ple (st_freshreg s7) (st_freshreg s10)
H18:Ple (st_freshstate s7) (st_freshstate s10)
H19:forall n : positive, (st_datapath s7) ! n = None \/ (st_datapath s10) ! n = (st_datapath s7) ! n
H20:forall n : positive, (st_controllogic s7) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s7) ! n
l0:Z.pos (max_pc_map (st_controllogic s10)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic s10))) Int.max_unsigned = left l0
l:Z.pos (max_pc_map (st_datapath s10)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath s10))) Int.max_unsigned = left l
H6:st_freshreg s2 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H5:st_freshreg s11 = Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))
H7:st_freshreg s0 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))
TR_INSTR:Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))) = st_freshreg s3
TR_DEC:st_freshreg s3 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H8:st_freshreg s5 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))
H9:st_freshreg s6 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))))
H10:st_freshreg s10 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))))
H4:st_controllogic s6 = st_controllogic s10
H11:st_controllogic s5 = st_controllogic s6
H12:st_controllogic s4 = st_controllogic s5
H13:st_controllogic s11 = st_controllogic s0
H14:st_controllogic s1 = st_controllogic s11
H15:st_controllogic s0 = st_controllogic s2
pc:positive
i:RTL.instruction
H101:(RTL.fn_code f) ! pc = Some i
EQ3D:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 (state_incr_intro s3 s4 e p p0 o o0)

tr_code (RTL.fn_code f) pc i (st_datapath s10) (st_controllogic s10) (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (st_st s10) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
f:RTL.function
s1, s10:state
H96:st_st s1 = st_st s10
H97:Ple (st_freshreg s1) (st_freshreg s10)
H98:Ple (st_freshstate s1) (st_freshstate s10)
H99:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s10) ! n = (st_datapath s1) ! n
H100:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s1) ! n
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s10
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s10)
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s10)
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s10) ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s10) ! n = (AssocMap.empty stmnt) ! n
s11:state
H91:st_st s1 = st_st s11
H92:Ple (st_freshreg s1) (st_freshreg s11)
H93:Ple (st_freshstate s1) (st_freshstate s11)
H94:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s11) ! n = (st_datapath s1) ! n
H95:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s11) ! n = (st_controllogic s1) ! n
H86:st_st s11 = st_st s10
H87:Ple (st_freshreg s11) (st_freshreg s10)
H88:Ple (st_freshstate s11) (st_freshstate s10)
H89:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s10) ! n = (st_datapath s11) ! n
H90:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s11) ! n
EQ:st_datapath s1 = st_datapath s11
s0:state
H81:st_st s11 = st_st s0
H82:Ple (st_freshreg s11) (st_freshreg s0)
H83:Ple (st_freshstate s11) (st_freshstate s0)
H84:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s0) ! n = (st_datapath s11) ! n
H85:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s11) ! n
H76:st_st s0 = st_st s10
H77:Ple (st_freshreg s0) (st_freshreg s10)
H78:Ple (st_freshstate s0) (st_freshstate s10)
H79:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s10) ! n = (st_datapath s0) ! n
H80:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s0) ! n
EQ1:st_datapath s11 = st_datapath s0
s2:state
H71:st_st s0 = st_st s2
H72:Ple (st_freshreg s0) (st_freshreg s2)
H73:Ple (st_freshstate s0) (st_freshstate s2)
H74:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
H75:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
H66:st_st s2 = st_st s10
H67:Ple (st_freshreg s2) (st_freshreg s10)
H68:Ple (st_freshstate s2) (st_freshstate s10)
H69:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s10) ! n = (st_datapath s2) ! n
H70:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s2) ! n
EQ0:st_datapath s0 = st_datapath s2
s3:state
e0:st_st s2 = st_st s3
p1:Ple (st_freshreg s2) (st_freshreg s3)
p2:Ple (st_freshstate s2) (st_freshstate s3)
o1:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s3) ! n = (st_datapath s2) ! n
o2:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s3) ! n = (st_controllogic s2) ! n
H61:st_st s3 = st_st s10
H62:Ple (st_freshreg s3) (st_freshreg s10)
H63:Ple (st_freshstate s3) (st_freshstate s10)
H64:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s10) ! n = (st_datapath s3) ! n
H65:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s3) ! n
EQ2:HTLMonadExtra.collectlist (transf_instr (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 (state_incr_intro s2 s3 e0 p1 p2 o1 o2)
s4:state
e:st_st s3 = st_st s4
p:Ple (st_freshreg s3) (st_freshreg s4)
p0:Ple (st_freshstate s3) (st_freshstate s4)
o:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s4) ! n = (st_datapath s3) ! n
o0:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s4) ! n = (st_controllogic s3) ! n
H56:st_st s4 = st_st s10
H57:Ple (st_freshreg s4) (st_freshreg s10)
H58:Ple (st_freshstate s4) (st_freshstate s10)
H59:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s10) ! n = (st_datapath s4) ! n
H60:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s4) ! n
EQ3:st_datapath s3 = st_datapath s4
s5:state
H51:st_st s4 = st_st s5
H52:Ple (st_freshreg s4) (st_freshreg s5)
H53:Ple (st_freshstate s4) (st_freshstate s5)
H54:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s5) ! n = (st_datapath s4) ! n
H55:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s5) ! n = (st_controllogic s4) ! n
H46:st_st s5 = st_st s10
H47:Ple (st_freshreg s5) (st_freshreg s10)
H48:Ple (st_freshstate s5) (st_freshstate s10)
H49:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s10) ! n = (st_datapath s5) ! n
H50:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s5) ! n
EQ4:st_datapath s4 = st_datapath s5
s6:state
H41:st_st s5 = st_st s6
H42:Ple (st_freshreg s5) (st_freshreg s6)
H43:Ple (st_freshstate s5) (st_freshstate s6)
H44:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s6) ! n = (st_datapath s5) ! n
H45:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s6) ! n = (st_controllogic s5) ! n
H36:st_st s6 = st_st s10
H37:Ple (st_freshreg s6) (st_freshreg s10)
H38:Ple (st_freshstate s6) (st_freshstate s10)
H39:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H40:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
EQ5:st_datapath s5 = st_datapath s6
H31:st_st s6 = st_st s10
H32:Ple (st_freshreg s6) (st_freshreg s10)
H33:Ple (st_freshstate s6) (st_freshstate s10)
H34:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H35:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
s9:state
H26:st_st s9 = st_st s10
H27:Ple (st_freshreg s9) (st_freshreg s10)
H28:Ple (st_freshstate s9) (st_freshstate s10)
H29:forall n : positive, (st_datapath s9) ! n = None \/ (st_datapath s10) ! n = (st_datapath s9) ! n
H30:forall n : positive, (st_controllogic s9) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s9) ! n
EQ6:st_datapath s6 = st_datapath s10
s8:state
H21:st_st s8 = st_st s10
H22:Ple (st_freshreg s8) (st_freshreg s10)
H23:Ple (st_freshstate s8) (st_freshstate s10)
H24:forall n : positive, (st_datapath s8) ! n = None \/ (st_datapath s10) ! n = (st_datapath s8) ! n
H25:forall n : positive, (st_controllogic s8) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s8) ! n
s7:state
H16:st_st s7 = st_st s10
H17:Ple (st_freshreg s7) (st_freshreg s10)
H18:Ple (st_freshstate s7) (st_freshstate s10)
H19:forall n : positive, (st_datapath s7) ! n = None \/ (st_datapath s10) ! n = (st_datapath s7) ! n
H20:forall n : positive, (st_controllogic s7) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s7) ! n
l0:Z.pos (max_pc_map (st_controllogic s10)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic s10))) Int.max_unsigned = left l0
l:Z.pos (max_pc_map (st_datapath s10)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath s10))) Int.max_unsigned = left l
H6:st_freshreg s2 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H5:st_freshreg s11 = Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))
H7:st_freshreg s0 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))
TR_INSTR:Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))) = st_freshreg s3
TR_DEC:st_freshreg s3 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H8:st_freshreg s5 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))
H9:st_freshreg s6 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))))
H10:st_freshreg s10 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))))
H4:st_controllogic s6 = st_controllogic s10
H11:st_controllogic s5 = st_controllogic s6
H12:st_controllogic s4 = st_controllogic s5
H13:st_controllogic s11 = st_controllogic s0
H14:st_controllogic s1 = st_controllogic s11
H15:st_controllogic s0 = st_controllogic s2
pc:positive
i:RTL.instruction
H101:(RTL.fn_code f) ! pc = Some i
EQ3D:HTLMonadExtra.collectlist (fun r : reg => declare_reg (Some Vinput) r 32) (RTL.fn_params f) s3 = OK tt s4 (state_incr_intro s3 s4 e p p0 o o0)

tr_code (RTL.fn_code f) pc i (st_datapath s10) (st_controllogic s10) (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (st_st s10) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
f:RTL.function
s1, s10:state
H96:st_st s1 = st_st s10
H97:Ple (st_freshreg s1) (st_freshreg s10)
H98:Ple (st_freshstate s1) (st_freshstate s10)
H99:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s10) ! n = (st_datapath s1) ! n
H100:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s1) ! n
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s10
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s10)
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s10)
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s10) ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s10) ! n = (AssocMap.empty stmnt) ! n
s11:state
H91:st_st s1 = st_st s11
H92:Ple (st_freshreg s1) (st_freshreg s11)
H93:Ple (st_freshstate s1) (st_freshstate s11)
H94:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s11) ! n = (st_datapath s1) ! n
H95:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s11) ! n = (st_controllogic s1) ! n
H86:st_st s11 = st_st s10
H87:Ple (st_freshreg s11) (st_freshreg s10)
H88:Ple (st_freshstate s11) (st_freshstate s10)
H89:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s10) ! n = (st_datapath s11) ! n
H90:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s11) ! n
EQ:st_datapath s1 = st_datapath s11
s0:state
H81:st_st s11 = st_st s0
H82:Ple (st_freshreg s11) (st_freshreg s0)
H83:Ple (st_freshstate s11) (st_freshstate s0)
H84:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s0) ! n = (st_datapath s11) ! n
H85:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s11) ! n
H76:st_st s0 = st_st s10
H77:Ple (st_freshreg s0) (st_freshreg s10)
H78:Ple (st_freshstate s0) (st_freshstate s10)
H79:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s10) ! n = (st_datapath s0) ! n
H80:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s0) ! n
EQ1:st_datapath s11 = st_datapath s0
s2:state
H71:st_st s0 = st_st s2
H72:Ple (st_freshreg s0) (st_freshreg s2)
H73:Ple (st_freshstate s0) (st_freshstate s2)
H74:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
H75:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
H66:st_st s2 = st_st s10
H67:Ple (st_freshreg s2) (st_freshreg s10)
H68:Ple (st_freshstate s2) (st_freshstate s10)
H69:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s10) ! n = (st_datapath s2) ! n
H70:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s2) ! n
EQ0:st_datapath s0 = st_datapath s2
s3:state
e0:st_st s2 = st_st s3
p1:Ple (st_freshreg s2) (st_freshreg s3)
p2:Ple (st_freshstate s2) (st_freshstate s3)
o1:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s3) ! n = (st_datapath s2) ! n
o2:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s3) ! n = (st_controllogic s2) ! n
H61:st_st s3 = st_st s10
H62:Ple (st_freshreg s3) (st_freshreg s10)
H63:Ple (st_freshstate s3) (st_freshstate s10)
H64:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s10) ! n = (st_datapath s3) ! n
H65:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s3) ! n
EQ2:HTLMonadExtra.collectlist (transf_instr (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 (state_incr_intro s2 s3 e0 p1 p2 o1 o2)
s4:state
e:st_st s3 = st_st s4
p:Ple (st_freshreg s3) (st_freshreg s4)
p0:Ple (st_freshstate s3) (st_freshstate s4)
o:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s4) ! n = (st_datapath s3) ! n
o0:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s4) ! n = (st_controllogic s3) ! n
H56:st_st s4 = st_st s10
H57:Ple (st_freshreg s4) (st_freshreg s10)
H58:Ple (st_freshstate s4) (st_freshstate s10)
H59:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s10) ! n = (st_datapath s4) ! n
H60:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s4) ! n
EQ3:st_datapath s3 = st_datapath s4
s5:state
H51:st_st s4 = st_st s5
H52:Ple (st_freshreg s4) (st_freshreg s5)
H53:Ple (st_freshstate s4) (st_freshstate s5)
H54:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s5) ! n = (st_datapath s4) ! n
H55:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s5) ! n = (st_controllogic s4) ! n
H46:st_st s5 = st_st s10
H47:Ple (st_freshreg s5) (st_freshreg s10)
H48:Ple (st_freshstate s5) (st_freshstate s10)
H49:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s10) ! n = (st_datapath s5) ! n
H50:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s5) ! n
EQ4:st_datapath s4 = st_datapath s5
s6:state
H41:st_st s5 = st_st s6
H42:Ple (st_freshreg s5) (st_freshreg s6)
H43:Ple (st_freshstate s5) (st_freshstate s6)
H44:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s6) ! n = (st_datapath s5) ! n
H45:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s6) ! n = (st_controllogic s5) ! n
H36:st_st s6 = st_st s10
H37:Ple (st_freshreg s6) (st_freshreg s10)
H38:Ple (st_freshstate s6) (st_freshstate s10)
H39:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H40:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
EQ5:st_datapath s5 = st_datapath s6
H31:st_st s6 = st_st s10
H32:Ple (st_freshreg s6) (st_freshreg s10)
H33:Ple (st_freshstate s6) (st_freshstate s10)
H34:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H35:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
s9:state
H26:st_st s9 = st_st s10
H27:Ple (st_freshreg s9) (st_freshreg s10)
H28:Ple (st_freshstate s9) (st_freshstate s10)
H29:forall n : positive, (st_datapath s9) ! n = None \/ (st_datapath s10) ! n = (st_datapath s9) ! n
H30:forall n : positive, (st_controllogic s9) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s9) ! n
EQ6:st_datapath s6 = st_datapath s10
s8:state
H21:st_st s8 = st_st s10
H22:Ple (st_freshreg s8) (st_freshreg s10)
H23:Ple (st_freshstate s8) (st_freshstate s10)
H24:forall n : positive, (st_datapath s8) ! n = None \/ (st_datapath s10) ! n = (st_datapath s8) ! n
H25:forall n : positive, (st_controllogic s8) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s8) ! n
s7:state
H16:st_st s7 = st_st s10
H17:Ple (st_freshreg s7) (st_freshreg s10)
H18:Ple (st_freshstate s7) (st_freshstate s10)
H19:forall n : positive, (st_datapath s7) ! n = None \/ (st_datapath s10) ! n = (st_datapath s7) ! n
H20:forall n : positive, (st_controllogic s7) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s7) ! n
l0:Z.pos (max_pc_map (st_controllogic s10)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic s10))) Int.max_unsigned = left l0
l:Z.pos (max_pc_map (st_datapath s10)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath s10))) Int.max_unsigned = left l
H6:st_freshreg s2 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H5:st_freshreg s11 = Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))
H7:st_freshreg s0 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))
TR_INSTR:Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))) = st_freshreg s3
TR_DEC:st_freshreg s3 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H8:st_freshreg s5 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))
H9:st_freshreg s6 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))))
H10:st_freshreg s10 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))))
H4:st_controllogic s6 = st_controllogic s10
H11:st_controllogic s5 = st_controllogic s6
H12:st_controllogic s4 = st_controllogic s5
H13:st_controllogic s11 = st_controllogic s0
H14:st_controllogic s1 = st_controllogic s11
H15:st_controllogic s0 = st_controllogic s2
pc:positive
i:RTL.instruction
H101:(RTL.fn_code f) ! pc = Some i
EQ3D:st_controllogic s3 = st_controllogic s4

tr_code (RTL.fn_code f) pc i (st_datapath s10) (st_controllogic s10) (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (st_st s10) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
f:RTL.function
s1, s10:state
H96:st_st s1 = st_st s10
H97:Ple (st_freshreg s1) (st_freshreg s10)
H98:Ple (st_freshstate s1) (st_freshstate s10)
H99:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s10) ! n = (st_datapath s1) ! n
H100:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s1) ! n
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s10
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s10)
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s10)
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s10) ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s10) ! n = (AssocMap.empty stmnt) ! n
s11:state
H91:st_st s1 = st_st s11
H92:Ple (st_freshreg s1) (st_freshreg s11)
H93:Ple (st_freshstate s1) (st_freshstate s11)
H94:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s11) ! n = (st_datapath s1) ! n
H95:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s11) ! n = (st_controllogic s1) ! n
H86:st_st s11 = st_st s10
H87:Ple (st_freshreg s11) (st_freshreg s10)
H88:Ple (st_freshstate s11) (st_freshstate s10)
H89:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s10) ! n = (st_datapath s11) ! n
H90:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s11) ! n
EQ:st_datapath s1 = st_datapath s11
s0:state
H81:st_st s11 = st_st s0
H82:Ple (st_freshreg s11) (st_freshreg s0)
H83:Ple (st_freshstate s11) (st_freshstate s0)
H84:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s0) ! n = (st_datapath s11) ! n
H85:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s11) ! n
H76:st_st s0 = st_st s10
H77:Ple (st_freshreg s0) (st_freshreg s10)
H78:Ple (st_freshstate s0) (st_freshstate s10)
H79:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s10) ! n = (st_datapath s0) ! n
H80:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s0) ! n
EQ1:st_datapath s11 = st_datapath s0
s2:state
H71:st_st s0 = st_st s2
H72:Ple (st_freshreg s0) (st_freshreg s2)
H73:Ple (st_freshstate s0) (st_freshstate s2)
H74:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
H75:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
H66:st_st s2 = st_st s10
H67:Ple (st_freshreg s2) (st_freshreg s10)
H68:Ple (st_freshstate s2) (st_freshstate s10)
H69:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s10) ! n = (st_datapath s2) ! n
H70:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s2) ! n
EQ0:st_datapath s0 = st_datapath s2
s3:state
e0:st_st s2 = st_st s3
p1:Ple (st_freshreg s2) (st_freshreg s3)
p2:Ple (st_freshstate s2) (st_freshstate s3)
o1:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s3) ! n = (st_datapath s2) ! n
o2:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s3) ! n = (st_controllogic s2) ! n
H61:st_st s3 = st_st s10
H62:Ple (st_freshreg s3) (st_freshreg s10)
H63:Ple (st_freshstate s3) (st_freshstate s10)
H64:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s10) ! n = (st_datapath s3) ! n
H65:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s3) ! n
EQ2:HTLMonadExtra.collectlist (transf_instr (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 (state_incr_intro s2 s3 e0 p1 p2 o1 o2)
s4:state
e:st_st s3 = st_st s4
p:Ple (st_freshreg s3) (st_freshreg s4)
p0:Ple (st_freshstate s3) (st_freshstate s4)
o:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s4) ! n = (st_datapath s3) ! n
o0:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s4) ! n = (st_controllogic s3) ! n
H56:st_st s4 = st_st s10
H57:Ple (st_freshreg s4) (st_freshreg s10)
H58:Ple (st_freshstate s4) (st_freshstate s10)
H59:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s10) ! n = (st_datapath s4) ! n
H60:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s4) ! n
EQ3:st_datapath s3 = st_datapath s4
s5:state
H51:st_st s4 = st_st s5
H52:Ple (st_freshreg s4) (st_freshreg s5)
H53:Ple (st_freshstate s4) (st_freshstate s5)
H54:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s5) ! n = (st_datapath s4) ! n
H55:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s5) ! n = (st_controllogic s4) ! n
H46:st_st s5 = st_st s10
H47:Ple (st_freshreg s5) (st_freshreg s10)
H48:Ple (st_freshstate s5) (st_freshstate s10)
H49:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s10) ! n = (st_datapath s5) ! n
H50:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s5) ! n
EQ4:st_datapath s4 = st_datapath s5
s6:state
H41:st_st s5 = st_st s6
H42:Ple (st_freshreg s5) (st_freshreg s6)
H43:Ple (st_freshstate s5) (st_freshstate s6)
H44:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s6) ! n = (st_datapath s5) ! n
H45:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s6) ! n = (st_controllogic s5) ! n
H36:st_st s6 = st_st s10
H37:Ple (st_freshreg s6) (st_freshreg s10)
H38:Ple (st_freshstate s6) (st_freshstate s10)
H39:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H40:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
EQ5:st_datapath s5 = st_datapath s6
H31:st_st s6 = st_st s10
H32:Ple (st_freshreg s6) (st_freshreg s10)
H33:Ple (st_freshstate s6) (st_freshstate s10)
H34:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H35:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
s9:state
H26:st_st s9 = st_st s10
H27:Ple (st_freshreg s9) (st_freshreg s10)
H28:Ple (st_freshstate s9) (st_freshstate s10)
H29:forall n : positive, (st_datapath s9) ! n = None \/ (st_datapath s10) ! n = (st_datapath s9) ! n
H30:forall n : positive, (st_controllogic s9) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s9) ! n
EQ6:st_datapath s6 = st_datapath s10
s8:state
H21:st_st s8 = st_st s10
H22:Ple (st_freshreg s8) (st_freshreg s10)
H23:Ple (st_freshstate s8) (st_freshstate s10)
H24:forall n : positive, (st_datapath s8) ! n = None \/ (st_datapath s10) ! n = (st_datapath s8) ! n
H25:forall n : positive, (st_controllogic s8) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s8) ! n
s7:state
H16:st_st s7 = st_st s10
H17:Ple (st_freshreg s7) (st_freshreg s10)
H18:Ple (st_freshstate s7) (st_freshstate s10)
H19:forall n : positive, (st_datapath s7) ! n = None \/ (st_datapath s10) ! n = (st_datapath s7) ! n
H20:forall n : positive, (st_controllogic s7) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s7) ! n
l0:Z.pos (max_pc_map (st_controllogic s10)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic s10))) Int.max_unsigned = left l0
l:Z.pos (max_pc_map (st_datapath s10)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath s10))) Int.max_unsigned = left l
H6:st_freshreg s2 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H5:st_freshreg s11 = Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))
H7:st_freshreg s0 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))
TR_INSTR:Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))) = st_freshreg s3
TR_DEC:st_freshreg s3 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H8:st_freshreg s5 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))
H9:st_freshreg s6 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))))
H10:st_freshreg s10 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))))
H4:st_controllogic s6 = st_controllogic s10
H11:st_controllogic s5 = st_controllogic s6
H12:st_controllogic s4 = st_controllogic s5
H13:st_controllogic s11 = st_controllogic s0
H14:st_controllogic s1 = st_controllogic s11
H15:st_controllogic s0 = st_controllogic s2
pc:positive
i:RTL.instruction
H101:(RTL.fn_code f) ! pc = Some i
EQ3D:st_controllogic s3 = st_controllogic s4

tr_code (RTL.fn_code f) pc i (st_datapath s10) (st_controllogic s3) (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (st_st s10) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
f:RTL.function
s1, s10:state
H96:st_st s1 = st_st s10
H97:Ple (st_freshreg s1) (st_freshreg s10)
H98:Ple (st_freshstate s1) (st_freshstate s10)
H99:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s10) ! n = (st_datapath s1) ! n
H100:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s1) ! n
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s10
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s10)
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s10)
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s10) ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s10) ! n = (AssocMap.empty stmnt) ! n
s11:state
H91:st_st s1 = st_st s11
H92:Ple (st_freshreg s1) (st_freshreg s11)
H93:Ple (st_freshstate s1) (st_freshstate s11)
H94:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s11) ! n = (st_datapath s1) ! n
H95:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s11) ! n = (st_controllogic s1) ! n
H86:st_st s11 = st_st s10
H87:Ple (st_freshreg s11) (st_freshreg s10)
H88:Ple (st_freshstate s11) (st_freshstate s10)
H89:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s10) ! n = (st_datapath s11) ! n
H90:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s11) ! n
EQ:st_datapath s1 = st_datapath s11
s0:state
H81:st_st s11 = st_st s0
H82:Ple (st_freshreg s11) (st_freshreg s0)
H83:Ple (st_freshstate s11) (st_freshstate s0)
H84:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s0) ! n = (st_datapath s11) ! n
H85:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s11) ! n
H76:st_st s0 = st_st s10
H77:Ple (st_freshreg s0) (st_freshreg s10)
H78:Ple (st_freshstate s0) (st_freshstate s10)
H79:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s10) ! n = (st_datapath s0) ! n
H80:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s0) ! n
EQ1:st_datapath s11 = st_datapath s0
s2:state
H71:st_st s0 = st_st s2
H72:Ple (st_freshreg s0) (st_freshreg s2)
H73:Ple (st_freshstate s0) (st_freshstate s2)
H74:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
H75:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
H66:st_st s2 = st_st s10
H67:Ple (st_freshreg s2) (st_freshreg s10)
H68:Ple (st_freshstate s2) (st_freshstate s10)
H69:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s10) ! n = (st_datapath s2) ! n
H70:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s2) ! n
EQ0:st_datapath s0 = st_datapath s2
s3:state
e0:st_st s2 = st_st s3
p1:Ple (st_freshreg s2) (st_freshreg s3)
p2:Ple (st_freshstate s2) (st_freshstate s3)
o1:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s3) ! n = (st_datapath s2) ! n
o2:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s3) ! n = (st_controllogic s2) ! n
H61:st_st s3 = st_st s10
H62:Ple (st_freshreg s3) (st_freshreg s10)
H63:Ple (st_freshstate s3) (st_freshstate s10)
H64:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s10) ! n = (st_datapath s3) ! n
H65:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s3) ! n
EQ2:HTLMonadExtra.collectlist (transf_instr (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 (state_incr_intro s2 s3 e0 p1 p2 o1 o2)
s4:state
e:st_st s3 = st_st s4
p:Ple (st_freshreg s3) (st_freshreg s4)
p0:Ple (st_freshstate s3) (st_freshstate s4)
o:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s4) ! n = (st_datapath s3) ! n
o0:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s4) ! n = (st_controllogic s3) ! n
H56:st_st s4 = st_st s10
H57:Ple (st_freshreg s4) (st_freshreg s10)
H58:Ple (st_freshstate s4) (st_freshstate s10)
H59:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s10) ! n = (st_datapath s4) ! n
H60:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s4) ! n
EQ3:st_datapath s3 = st_datapath s4
s5:state
H51:st_st s4 = st_st s5
H52:Ple (st_freshreg s4) (st_freshreg s5)
H53:Ple (st_freshstate s4) (st_freshstate s5)
H54:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s5) ! n = (st_datapath s4) ! n
H55:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s5) ! n = (st_controllogic s4) ! n
H46:st_st s5 = st_st s10
H47:Ple (st_freshreg s5) (st_freshreg s10)
H48:Ple (st_freshstate s5) (st_freshstate s10)
H49:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s10) ! n = (st_datapath s5) ! n
H50:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s5) ! n
EQ4:st_datapath s4 = st_datapath s5
s6:state
H41:st_st s5 = st_st s6
H42:Ple (st_freshreg s5) (st_freshreg s6)
H43:Ple (st_freshstate s5) (st_freshstate s6)
H44:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s6) ! n = (st_datapath s5) ! n
H45:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s6) ! n = (st_controllogic s5) ! n
H36:st_st s6 = st_st s10
H37:Ple (st_freshreg s6) (st_freshreg s10)
H38:Ple (st_freshstate s6) (st_freshstate s10)
H39:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H40:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
EQ5:st_datapath s5 = st_datapath s6
H31:st_st s6 = st_st s10
H32:Ple (st_freshreg s6) (st_freshreg s10)
H33:Ple (st_freshstate s6) (st_freshstate s10)
H34:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H35:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
s9:state
H26:st_st s9 = st_st s10
H27:Ple (st_freshreg s9) (st_freshreg s10)
H28:Ple (st_freshstate s9) (st_freshstate s10)
H29:forall n : positive, (st_datapath s9) ! n = None \/ (st_datapath s10) ! n = (st_datapath s9) ! n
H30:forall n : positive, (st_controllogic s9) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s9) ! n
EQ6:st_datapath s6 = st_datapath s10
s8:state
H21:st_st s8 = st_st s10
H22:Ple (st_freshreg s8) (st_freshreg s10)
H23:Ple (st_freshstate s8) (st_freshstate s10)
H24:forall n : positive, (st_datapath s8) ! n = None \/ (st_datapath s10) ! n = (st_datapath s8) ! n
H25:forall n : positive, (st_controllogic s8) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s8) ! n
s7:state
H16:st_st s7 = st_st s10
H17:Ple (st_freshreg s7) (st_freshreg s10)
H18:Ple (st_freshstate s7) (st_freshstate s10)
H19:forall n : positive, (st_datapath s7) ! n = None \/ (st_datapath s10) ! n = (st_datapath s7) ! n
H20:forall n : positive, (st_controllogic s7) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s7) ! n
l0:Z.pos (max_pc_map (st_controllogic s10)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic s10))) Int.max_unsigned = left l0
l:Z.pos (max_pc_map (st_datapath s10)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath s10))) Int.max_unsigned = left l
H6:st_freshreg s2 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H5:st_freshreg s11 = Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))
H7:st_freshreg s0 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))
TR_INSTR:Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))) = st_freshreg s3
TR_DEC:st_freshreg s3 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H8:st_freshreg s5 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))
H9:st_freshreg s6 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))))
H10:st_freshreg s10 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))))
H4:st_controllogic s6 = st_controllogic s10
H11:st_controllogic s5 = st_controllogic s6
H12:st_controllogic s4 = st_controllogic s5
H13:st_controllogic s11 = st_controllogic s0
H14:st_controllogic s1 = st_controllogic s11
H15:st_controllogic s0 = st_controllogic s2
pc:positive
i:RTL.instruction
H101:(RTL.fn_code f) ! pc = Some i
EQ3D:st_controllogic s3 = st_controllogic s4

tr_code (RTL.fn_code f) pc i (st_datapath s3) (st_controllogic s3) (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (st_st s10) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
f:RTL.function
s1, s10:state
H96:st_st s1 = st_st s10
H97:Ple (st_freshreg s1) (st_freshreg s10)
H98:Ple (st_freshstate s1) (st_freshstate s10)
H99:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s10) ! n = (st_datapath s1) ! n
H100:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s1) ! n
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s10
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s10)
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s10)
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s10) ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s10) ! n = (AssocMap.empty stmnt) ! n
s11:state
H91:st_st s1 = st_st s11
H92:Ple (st_freshreg s1) (st_freshreg s11)
H93:Ple (st_freshstate s1) (st_freshstate s11)
H94:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s11) ! n = (st_datapath s1) ! n
H95:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s11) ! n = (st_controllogic s1) ! n
H86:st_st s11 = st_st s10
H87:Ple (st_freshreg s11) (st_freshreg s10)
H88:Ple (st_freshstate s11) (st_freshstate s10)
H89:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s10) ! n = (st_datapath s11) ! n
H90:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s11) ! n
EQ:st_datapath s1 = st_datapath s11
s0:state
H81:st_st s11 = st_st s0
H82:Ple (st_freshreg s11) (st_freshreg s0)
H83:Ple (st_freshstate s11) (st_freshstate s0)
H84:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s0) ! n = (st_datapath s11) ! n
H85:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s11) ! n
H76:st_st s0 = st_st s10
H77:Ple (st_freshreg s0) (st_freshreg s10)
H78:Ple (st_freshstate s0) (st_freshstate s10)
H79:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s10) ! n = (st_datapath s0) ! n
H80:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s0) ! n
EQ1:st_datapath s11 = st_datapath s0
s2:state
H71:st_st s0 = st_st s2
H72:Ple (st_freshreg s0) (st_freshreg s2)
H73:Ple (st_freshstate s0) (st_freshstate s2)
H74:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
H75:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
H66:st_st s2 = st_st s10
H67:Ple (st_freshreg s2) (st_freshreg s10)
H68:Ple (st_freshstate s2) (st_freshstate s10)
H69:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s10) ! n = (st_datapath s2) ! n
H70:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s2) ! n
EQ0:st_datapath s0 = st_datapath s2
s3:state
e0:st_st s2 = st_st s3
p1:Ple (st_freshreg s2) (st_freshreg s3)
p2:Ple (st_freshstate s2) (st_freshstate s3)
o1:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s3) ! n = (st_datapath s2) ! n
o2:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s3) ! n = (st_controllogic s2) ! n
H61:st_st s3 = st_st s10
H62:Ple (st_freshreg s3) (st_freshreg s10)
H63:Ple (st_freshstate s3) (st_freshstate s10)
H64:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s10) ! n = (st_datapath s3) ! n
H65:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s3) ! n
EQ2:HTLMonadExtra.collectlist (transf_instr (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 (state_incr_intro s2 s3 e0 p1 p2 o1 o2)
s4:state
e:st_st s3 = st_st s4
p:Ple (st_freshreg s3) (st_freshreg s4)
p0:Ple (st_freshstate s3) (st_freshstate s4)
o:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s4) ! n = (st_datapath s3) ! n
o0:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s4) ! n = (st_controllogic s3) ! n
H56:st_st s4 = st_st s10
H57:Ple (st_freshreg s4) (st_freshreg s10)
H58:Ple (st_freshstate s4) (st_freshstate s10)
H59:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s10) ! n = (st_datapath s4) ! n
H60:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s4) ! n
EQ3:st_datapath s3 = st_datapath s4
s5:state
H51:st_st s4 = st_st s5
H52:Ple (st_freshreg s4) (st_freshreg s5)
H53:Ple (st_freshstate s4) (st_freshstate s5)
H54:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s5) ! n = (st_datapath s4) ! n
H55:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s5) ! n = (st_controllogic s4) ! n
H46:st_st s5 = st_st s10
H47:Ple (st_freshreg s5) (st_freshreg s10)
H48:Ple (st_freshstate s5) (st_freshstate s10)
H49:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s10) ! n = (st_datapath s5) ! n
H50:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s5) ! n
EQ4:st_datapath s4 = st_datapath s5
s6:state
H41:st_st s5 = st_st s6
H42:Ple (st_freshreg s5) (st_freshreg s6)
H43:Ple (st_freshstate s5) (st_freshstate s6)
H44:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s6) ! n = (st_datapath s5) ! n
H45:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s6) ! n = (st_controllogic s5) ! n
H36:st_st s6 = st_st s10
H37:Ple (st_freshreg s6) (st_freshreg s10)
H38:Ple (st_freshstate s6) (st_freshstate s10)
H39:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H40:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
EQ5:st_datapath s5 = st_datapath s6
H31:st_st s6 = st_st s10
H32:Ple (st_freshreg s6) (st_freshreg s10)
H33:Ple (st_freshstate s6) (st_freshstate s10)
H34:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H35:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
s9:state
H26:st_st s9 = st_st s10
H27:Ple (st_freshreg s9) (st_freshreg s10)
H28:Ple (st_freshstate s9) (st_freshstate s10)
H29:forall n : positive, (st_datapath s9) ! n = None \/ (st_datapath s10) ! n = (st_datapath s9) ! n
H30:forall n : positive, (st_controllogic s9) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s9) ! n
EQ6:st_datapath s6 = st_datapath s10
s8:state
H21:st_st s8 = st_st s10
H22:Ple (st_freshreg s8) (st_freshreg s10)
H23:Ple (st_freshstate s8) (st_freshstate s10)
H24:forall n : positive, (st_datapath s8) ! n = None \/ (st_datapath s10) ! n = (st_datapath s8) ! n
H25:forall n : positive, (st_controllogic s8) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s8) ! n
s7:state
H16:st_st s7 = st_st s10
H17:Ple (st_freshreg s7) (st_freshreg s10)
H18:Ple (st_freshstate s7) (st_freshstate s10)
H19:forall n : positive, (st_datapath s7) ! n = None \/ (st_datapath s10) ! n = (st_datapath s7) ! n
H20:forall n : positive, (st_controllogic s7) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s7) ! n
l0:Z.pos (max_pc_map (st_controllogic s10)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic s10))) Int.max_unsigned = left l0
l:Z.pos (max_pc_map (st_datapath s10)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath s10))) Int.max_unsigned = left l
H6:st_freshreg s2 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H5:st_freshreg s11 = Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))
H7:st_freshreg s0 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))
TR_INSTR:Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))) = st_freshreg s3
TR_DEC:st_freshreg s3 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H8:st_freshreg s5 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))
H9:st_freshreg s6 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))))
H10:st_freshreg s10 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))))
H4:st_controllogic s6 = st_controllogic s10
H11:st_controllogic s5 = st_controllogic s6
H12:st_controllogic s4 = st_controllogic s5
H13:st_controllogic s11 = st_controllogic s0
H14:st_controllogic s1 = st_controllogic s11
H15:st_controllogic s0 = st_controllogic s2
pc:positive
i:RTL.instruction
H101:(RTL.fn_code f) ! pc = Some i
EQ3D:st_controllogic s3 = st_controllogic s4

tr_code (RTL.fn_code f) pc i (st_datapath s3) (st_controllogic s3) (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (st_st s3) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
f:RTL.function
s1, s10:state
H96:st_st s1 = st_st s10
H97:Ple (st_freshreg s1) (st_freshreg s10)
H98:Ple (st_freshstate s1) (st_freshstate s10)
H99:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s10) ! n = (st_datapath s1) ! n
H100:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s1) ! n
STACK_BOUND_LOW:0 <= RTL.fn_stacksize f
STACK_BOUND_HIGH:RTL.fn_stacksize f < 4294967296
STACK_ALIGN:RTL.fn_stacksize f mod 4 = 0
H:Pos.succ (RTL.max_reg_function f) = st_st s10
H0:Ple (Pos.succ (Pos.succ (RTL.max_reg_function f))) (st_freshreg s10)
H1:Ple (Pos.succ (RTL.max_pc_function f)) (st_freshstate s10)
H2:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_datapath s10) ! n = (AssocMap.empty stmnt) ! n
H3:forall n : positive, (AssocMap.empty stmnt) ! n = None \/ (st_controllogic s10) ! n = (AssocMap.empty stmnt) ! n
s11:state
H91:st_st s1 = st_st s11
H92:Ple (st_freshreg s1) (st_freshreg s11)
H93:Ple (st_freshstate s1) (st_freshstate s11)
H94:forall n : positive, (st_datapath s1) ! n = None \/ (st_datapath s11) ! n = (st_datapath s1) ! n
H95:forall n : positive, (st_controllogic s1) ! n = None \/ (st_controllogic s11) ! n = (st_controllogic s1) ! n
H86:st_st s11 = st_st s10
H87:Ple (st_freshreg s11) (st_freshreg s10)
H88:Ple (st_freshstate s11) (st_freshstate s10)
H89:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s10) ! n = (st_datapath s11) ! n
H90:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s11) ! n
EQ:st_datapath s1 = st_datapath s11
s0:state
H81:st_st s11 = st_st s0
H82:Ple (st_freshreg s11) (st_freshreg s0)
H83:Ple (st_freshstate s11) (st_freshstate s0)
H84:forall n : positive, (st_datapath s11) ! n = None \/ (st_datapath s0) ! n = (st_datapath s11) ! n
H85:forall n : positive, (st_controllogic s11) ! n = None \/ (st_controllogic s0) ! n = (st_controllogic s11) ! n
H76:st_st s0 = st_st s10
H77:Ple (st_freshreg s0) (st_freshreg s10)
H78:Ple (st_freshstate s0) (st_freshstate s10)
H79:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s10) ! n = (st_datapath s0) ! n
H80:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s0) ! n
EQ1:st_datapath s11 = st_datapath s0
s2:state
H71:st_st s0 = st_st s2
H72:Ple (st_freshreg s0) (st_freshreg s2)
H73:Ple (st_freshstate s0) (st_freshstate s2)
H74:forall n : positive, (st_datapath s0) ! n = None \/ (st_datapath s2) ! n = (st_datapath s0) ! n
H75:forall n : positive, (st_controllogic s0) ! n = None \/ (st_controllogic s2) ! n = (st_controllogic s0) ! n
H66:st_st s2 = st_st s10
H67:Ple (st_freshreg s2) (st_freshreg s10)
H68:Ple (st_freshstate s2) (st_freshstate s10)
H69:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s10) ! n = (st_datapath s2) ! n
H70:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s2) ! n
EQ0:st_datapath s0 = st_datapath s2
s3:state
e0:st_st s2 = st_st s3
p1:Ple (st_freshreg s2) (st_freshreg s3)
p2:Ple (st_freshstate s2) (st_freshstate s3)
o1:forall n : positive, (st_datapath s2) ! n = None \/ (st_datapath s3) ! n = (st_datapath s2) ! n
o2:forall n : positive, (st_controllogic s2) ! n = None \/ (st_controllogic s3) ! n = (st_controllogic s2) ! n
H61:st_st s3 = st_st s10
H62:Ple (st_freshreg s3) (st_freshreg s10)
H63:Ple (st_freshstate s3) (st_freshstate s10)
H64:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s10) ! n = (st_datapath s3) ! n
H65:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s3) ! n
EQ2:HTLMonadExtra.collectlist (transf_instr (Pos.succ (Pos.succ (RTL.max_reg_function f))) (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))) (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))) (PTree.elements (RTL.fn_code f)) s2 = OK tt s3 (state_incr_intro s2 s3 e0 p1 p2 o1 o2)
s4:state
e:st_st s3 = st_st s4
p:Ple (st_freshreg s3) (st_freshreg s4)
p0:Ple (st_freshstate s3) (st_freshstate s4)
o:forall n : positive, (st_datapath s3) ! n = None \/ (st_datapath s4) ! n = (st_datapath s3) ! n
o0:forall n : positive, (st_controllogic s3) ! n = None \/ (st_controllogic s4) ! n = (st_controllogic s3) ! n
H56:st_st s4 = st_st s10
H57:Ple (st_freshreg s4) (st_freshreg s10)
H58:Ple (st_freshstate s4) (st_freshstate s10)
H59:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s10) ! n = (st_datapath s4) ! n
H60:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s4) ! n
EQ3:st_datapath s3 = st_datapath s4
s5:state
H51:st_st s4 = st_st s5
H52:Ple (st_freshreg s4) (st_freshreg s5)
H53:Ple (st_freshstate s4) (st_freshstate s5)
H54:forall n : positive, (st_datapath s4) ! n = None \/ (st_datapath s5) ! n = (st_datapath s4) ! n
H55:forall n : positive, (st_controllogic s4) ! n = None \/ (st_controllogic s5) ! n = (st_controllogic s4) ! n
H46:st_st s5 = st_st s10
H47:Ple (st_freshreg s5) (st_freshreg s10)
H48:Ple (st_freshstate s5) (st_freshstate s10)
H49:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s10) ! n = (st_datapath s5) ! n
H50:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s5) ! n
EQ4:st_datapath s4 = st_datapath s5
s6:state
H41:st_st s5 = st_st s6
H42:Ple (st_freshreg s5) (st_freshreg s6)
H43:Ple (st_freshstate s5) (st_freshstate s6)
H44:forall n : positive, (st_datapath s5) ! n = None \/ (st_datapath s6) ! n = (st_datapath s5) ! n
H45:forall n : positive, (st_controllogic s5) ! n = None \/ (st_controllogic s6) ! n = (st_controllogic s5) ! n
H36:st_st s6 = st_st s10
H37:Ple (st_freshreg s6) (st_freshreg s10)
H38:Ple (st_freshstate s6) (st_freshstate s10)
H39:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H40:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
EQ5:st_datapath s5 = st_datapath s6
H31:st_st s6 = st_st s10
H32:Ple (st_freshreg s6) (st_freshreg s10)
H33:Ple (st_freshstate s6) (st_freshstate s10)
H34:forall n : positive, (st_datapath s6) ! n = None \/ (st_datapath s10) ! n = (st_datapath s6) ! n
H35:forall n : positive, (st_controllogic s6) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s6) ! n
s9:state
H26:st_st s9 = st_st s10
H27:Ple (st_freshreg s9) (st_freshreg s10)
H28:Ple (st_freshstate s9) (st_freshstate s10)
H29:forall n : positive, (st_datapath s9) ! n = None \/ (st_datapath s10) ! n = (st_datapath s9) ! n
H30:forall n : positive, (st_controllogic s9) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s9) ! n
EQ6:st_datapath s6 = st_datapath s10
s8:state
H21:st_st s8 = st_st s10
H22:Ple (st_freshreg s8) (st_freshreg s10)
H23:Ple (st_freshstate s8) (st_freshstate s10)
H24:forall n : positive, (st_datapath s8) ! n = None \/ (st_datapath s10) ! n = (st_datapath s8) ! n
H25:forall n : positive, (st_controllogic s8) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s8) ! n
s7:state
H16:st_st s7 = st_st s10
H17:Ple (st_freshreg s7) (st_freshreg s10)
H18:Ple (st_freshstate s7) (st_freshstate s10)
H19:forall n : positive, (st_datapath s7) ! n = None \/ (st_datapath s10) ! n = (st_datapath s7) ! n
H20:forall n : positive, (st_controllogic s7) ! n = None \/ (st_controllogic s10) ! n = (st_controllogic s7) ! n
l0:Z.pos (max_pc_map (st_controllogic s10)) <= Int.max_unsigned
Heqs0:zle (Z.pos (max_pc_map (st_controllogic s10))) Int.max_unsigned = left l0
l:Z.pos (max_pc_map (st_datapath s10)) <= Int.max_unsigned
Heqs9:zle (Z.pos (max_pc_map (st_datapath s10))) Int.max_unsigned = left l
H6:st_freshreg s2 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H5:st_freshreg s11 = Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))
H7:st_freshreg s0 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))
TR_INSTR:Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))) = st_freshreg s3
TR_DEC:st_freshreg s3 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))
H8:st_freshreg s5 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))
H9:st_freshreg s6 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f)))))))
H10:st_freshreg s10 = Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (Pos.succ (RTL.max_reg_function f))))))))
H4:st_controllogic s6 = st_controllogic s10
H11:st_controllogic s5 = st_controllogic s6
H12:st_controllogic s4 = st_controllogic s5
H13:st_controllogic s11 = st_controllogic s0
H14:st_controllogic s1 = st_controllogic s11
H15:st_controllogic s0 = st_controllogic s2
pc:positive
i:RTL.instruction
H101:(RTL.fn_code f) ! pc = Some i
EQ3D:st_controllogic s3 = st_controllogic s4

forall (pc : node) (instr : RTL.instruction), In (pc, instr) (PTree.elements (RTL.fn_code f)) -> (RTL.fn_code f) ! pc = Some instr
apply PTree.elements_complete. Qed.