Gible Seq Generation Proof

Defining a find block specification

Basically, it should be able to find the location of the block without using the find_block function, so that this is more useful for the proofs. There are various different types of options that could come up though:

  1. The instruction is a standard instruction present inside of a basic block.

  2. The instruction is a standard instruction which ends with a goto.

  3. The instruction is a control-flow instruction.

For case number 1, there should exist a value in the list of instructions, such that the instructions match exactly, and the indices match as well. In the original code, this instruction must have been going from the current node to the node - 1.

For case number 2, there should be an instruction at the right index again, however, this time there will also be a goto instruction in the control-flow part of the basic block.

For case number 3, there should be a nop instruction in the basic block, and then the equivalent control-flow instruction ending the basic block.

In the end though, it seems like two cases are actually enough, as the last two cases are similar enough that they can be merged into one.

Definition all_max {A} (c: PTree.t A) p: Prop :=
  Forall (fun x => x <= p) (map fst (PTree.elements c)).

Definition offset (pc pc': positive): nat := Pos.to_nat pc' - Pos.to_nat pc.

Section CORRECTNESS.

  Context (prog : RTL.program).
  Context (tprog : GibleSeq.program).

  Let ge : RTL.genv := Globalenvs.Genv.globalenv prog.
  Let tge : genv := Globalenvs.Genv.globalenv tprog.

Matches the basic block that should be present in the state. This simulates the small step execution of the basic block from the big step semantics that are currently present.

Why does it not need to find the pc’ value using find_block?

It doesn’t have to find the value because it’s given as an input, and the finding is actually done at that higher level already.

(*  Variant match_bblock (tc: code) (pc pc': node): list instr -> Prop :=
    | match_bblock_intro :
      forall bb cf,
        tc ! pc' = Some (mk_bblock bb cf) ->
        match_bblock tc pc pc' (list_drop (offset pc pc') bb).*)

  Definition imm_succ (pc pc': node) : Prop := pc' = Pos.pred pc.

  Definition valid_succ (tc: code) (pc: node) : Prop := exists b, tc ! pc = Some b.

  Inductive match_block (c: RTL.code) (tc: code) (pc: node): SeqBB.t -> Prop :=
  (*
   * Basic Block Instructions
   *)
  | match_block_nop b pc':
    c ! pc = Some (RTL.Inop pc') ->
    match_block c tc pc' b ->
    match_block c tc pc (RBnop :: b)
  | match_block_op b op args dst pc':
    c ! pc = Some (RTL.Iop op args dst pc') ->
    match_block c tc pc' b ->
    match_block c tc pc (RBop None op args dst :: b)
  | match_block_load b chunk addr args dst pc':
    c ! pc = Some (RTL.Iload chunk addr args dst pc') ->
    match_block c tc pc' b ->
    match_block c tc pc (RBload None chunk addr args dst :: b)
  | match_block_store b chunk addr args src pc':
    c ! pc = Some (RTL.Istore chunk addr args src pc') ->
    match_block c tc pc' b ->
    match_block c tc pc (RBstore None chunk addr args src :: b)
  (*
   * Control flow instructions using goto
   *)
  | match_block_goto pc':
    c ! pc = Some (RTL.Inop pc') ->
    valid_succ tc pc' ->
    match_block c tc pc (RBnop :: RBexit None (RBgoto pc') :: nil)
  | match_block_op_cf pc' op args dst:
    c ! pc = Some (RTL.Iop op args dst pc') ->
    valid_succ tc pc' ->
    match_block c tc pc (RBop None op args dst :: RBexit None (RBgoto pc') :: nil)
  | match_block_load_cf pc' chunk addr args dst:
    c ! pc = Some (RTL.Iload chunk addr args dst pc') ->
    valid_succ tc pc' ->
    match_block c tc pc (RBload None chunk addr args dst :: RBexit None (RBgoto pc') :: nil)
  | match_block_store_cf pc' chunk addr args src:
    c ! pc = Some (RTL.Istore chunk addr args src pc') ->
    valid_succ tc pc' ->
    match_block c tc pc (RBstore None chunk addr args src :: RBexit None (RBgoto pc') :: nil)
  (*
   * Standard control flow instructions
   *)
  | match_block_call sig ident args dst pc' :
    c ! pc = Some (RTL.Icall sig ident args dst pc') ->
    valid_succ tc pc' ->
    match_block c tc pc (RBnop :: RBexit None (RBcall sig ident args dst pc') :: nil)
  | match_block_tailcall sig ident args :
    c ! pc = Some (RTL.Itailcall sig ident args) ->
    match_block c tc pc (RBnop :: RBexit None (RBtailcall sig ident args) :: nil)
  | match_block_builtin ident args dst pc' :
    c ! pc = Some (RTL.Ibuiltin ident args dst pc') ->
    valid_succ tc pc' ->
    match_block c tc pc (RBnop :: RBexit None (RBbuiltin ident args dst pc') :: nil)
  | match_block_cond cond args pct pcf :
    c ! pc = Some (RTL.Icond cond args pct pcf) ->
    valid_succ tc pct ->
    valid_succ tc pcf ->
    match_block c tc pc (RBnop :: RBexit None (RBcond cond args pct pcf) :: nil)
  | match_block_jumptable r ns :
    c ! pc = Some (RTL.Ijumptable r ns) ->
    Forall (valid_succ tc) ns ->
    match_block c tc pc (RBnop :: RBexit None (RBjumptable r ns) :: nil)
  | match_block_return r :
    c ! pc = Some (RTL.Ireturn r) ->
    match_block c tc pc (RBnop :: RBexit None (RBreturn r) :: nil)
  .

Match the code

The match_code predicate asserts that for any node in the original control-flow graph, there is now a basic block in the new control- and data-flow graph that contains the same instruction, but also that the whole basic block matches some sequence of instructions starting at the node that corresponds to the basic block.

  Definition match_code (c: RTL.code) (tc: code) : Prop :=
    forall pc b, tc ! pc = Some b -> match_block c tc pc b.

  Variant match_stackframe : RTL.stackframe -> stackframe -> Prop :=
    | match_stackframe_init :
      forall res f tf sp pc rs
        (TF: transl_function f = OK tf)
        (VALID: valid_succ tf.(fn_code) pc),
        match_stackframe (RTL.Stackframe res f sp pc rs)
                         (Stackframe res tf sp pc rs (PMap.init false)).

  Definition sem_extrap f pc sp in_s in_s' block :=
    forall out_s block2,
      SeqBB.step tge sp in_s block out_s ->
      f.(fn_code) ! pc = Some block2 ->
      SeqBB.step tge sp in_s' block2 out_s.

  
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv

forall (c : RTL.code) (tc : code) (pc : node) (a : SeqBB.t), match_block c tc pc a -> exists i : RTL.instruction, c ! pc = Some i
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv

forall (c : RTL.code) (tc : code) (pc : node) (a : SeqBB.t), match_block c tc pc a -> exists i : RTL.instruction, c ! pc = Some i
inversion 1; eexists; eauto. Qed.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv

forall (c : RTL.code) (tc : code) (pc : node) (a : SeqBB.t), match_block c tc pc a -> a <> nil
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv

forall (c : RTL.code) (tc : code) (pc : node) (a : SeqBB.t), match_block c tc pc a -> a <> nil
inversion 1; crush. Qed.

Matching states

Initially, the idea was to define the match_states predicate normally to defines how to find the correct bb that should be executed, as well as the value of pc. However, this does not quite work when proving the equivalence of the translation from RTL to RTLBlock, because one cannot match one transition to a transition in RTLBlock. The alternative to this is to include a proof inside of the match_states that shows that the execution from the pc of the start of the basic block to the current point is the same as the whole execution (in one big step) of the basic block.

  Variant match_states : option SeqBB.t -> RTL.state -> state -> Prop :=
    | match_state :
      forall stk stk' f tf sp pc rs m pc0 b rs0 m0
        (TF: transl_function f = OK tf)
        (* TODO: I can remove the following [match_code]. *)
        (CODE: match_code f.(RTL.fn_code) tf.(fn_code))
        (BLOCK: exists b', tf.(fn_code) ! pc0 = Some b'
                           /\ match_block f.(RTL.fn_code) tf.(fn_code) pc b)
        (STK: Forall2 match_stackframe stk stk')
        (STARSIMU: star RTL.step ge (RTL.State stk f sp pc0 rs0 m0)
                                  E0 (RTL.State stk f sp pc rs m))
        (BB: sem_extrap tf pc0 sp (Iexec (mk_instr_state rs (PMap.init false) m))
                        (Iexec (mk_instr_state rs0 (PMap.init false) m0)) b),
        match_states (Some b) (RTL.State stk f sp pc rs m)
                     (State stk' tf sp pc0 rs0 (PMap.init false) m0)
    | match_callstate :
      forall cs cs' f tf args m
        (TF: transl_fundef f = OK tf)
        (STK: Forall2 match_stackframe cs cs'),
        match_states None (RTL.Callstate cs f args m) (Callstate cs' tf args m)
    | match_returnstate :
      forall cs cs' v m
        (STK: Forall2 match_stackframe cs cs'),
        match_states None (RTL.Returnstate cs v m) (Returnstate cs' v m)
  .

  Definition match_prog (p: RTL.program) (tp: GibleSeq.program) :=
    Linking.match_program (fun cu f tf => transl_fundef f = Errors.OK tf) eq p tp.

  
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv

forall (p : RTL.program) (tp : program), transl_program p = OK tp -> match_prog p tp
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv

forall (p : RTL.program) (tp : program), transl_program p = OK tp -> match_prog p tp
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
p: RTL.program
tp: program
H: transl_program p = OK tp

match_prog p tp
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
p: RTL.program
tp: program
H: transl_program p = OK tp

Linking.match_program (fun (_ : AST.program (AST.fundef RTL.function) unit) (f : AST.fundef RTL.function) (tf : AST.fundef function) => transl_fundef f = OK tf) eq p tp
eapply Linking.match_transform_partial_program; auto. Qed. Context (TRANSL : match_prog prog tprog).
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall s : ident, Genv.find_symbol tge s = Genv.find_symbol ge s
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall s : ident, Genv.find_symbol tge s = Genv.find_symbol ge s
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: ident

Genv.find_symbol tge s = Genv.find_symbol ge s
eapply (Genv.find_symbol_match TRANSL). Qed.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

Senv.equiv ge tge
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

Senv.equiv ge tge
intros; eapply (Genv.senv_transf_partial TRANSL). Qed. #[local] Hint Resolve senv_preserved : rtlbg.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (b : block) (f : RTL.fundef), Genv.find_funct_ptr ge b = Some f -> exists tf : fundef, Genv.find_funct_ptr tge b = Some tf /\ transl_fundef f = OK tf
Proof (Genv.find_funct_ptr_transf_partial TRANSL).
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (f : RTL.fundef) (tf : fundef), transl_fundef f = OK tf -> funsig tf = RTL.funsig f
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (f : RTL.fundef) (tf : fundef), transl_fundef f = OK tf -> funsig tf = RTL.funsig f
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (f : RTL.fundef) (tf : fundef), transf_partial_fundef transl_function f = OK tf -> funsig tf = RTL.funsig f
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (f : RTL.fundef) (tf : fundef), match f with | Internal f0 => bind (transl_function f0) (fun f' : function => OK (Internal f')) | External ef => OK (External ef) end = OK tf -> funsig tf = RTL.funsig f
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.fundef
tf: fundef
H: match f with | Internal f => bind (transl_function f) (fun f' : function => OK (Internal f')) | External ef => OK (External ef) end = OK tf

funsig tf = RTL.funsig f
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.fundef
tf: fundef
f0: RTL.function
Heqf0: f = Internal f0
H: bind (transl_function f0) (fun f' : function => OK (Internal f')) = OK tf

funsig tf = RTL.funsig (Internal f0)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.fundef
tf: fundef
e: external_function
Heqf0: f = External e
H: OK (External e) = OK tf
funsig tf = RTL.funsig (External e)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.fundef
tf: fundef
f0: RTL.function
Heqf0: f = Internal f0
H: match transl_function f0 with | OK x => OK (Internal x) | Error msg => Error msg end = OK tf

funsig tf = RTL.funsig (Internal f0)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.fundef
tf: fundef
e: external_function
Heqf0: f = External e
H: OK (External e) = OK tf
funsig tf = RTL.funsig (External e)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.fundef
tf: fundef
f0: RTL.function
Heqf0: f = Internal f0
f1: function
Heqr: transl_function f0 = OK f1
H: OK (Internal f1) = OK tf

funsig tf = RTL.funsig (Internal f0)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.fundef
tf: fundef
e: external_function
Heqf0: f = External e
H: OK (External e) = OK tf
funsig tf = RTL.funsig (External e)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f0: RTL.function
f1: function
Heqr: transl_function f0 = OK f1

funsig (Internal f1) = RTL.funsig (Internal f0)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.fundef
tf: fundef
e: external_function
Heqf0: f = External e
H: OK (External e) = OK tf
funsig tf = RTL.funsig (External e)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f0: RTL.function
f1: function
Heqr: match partition f0 with | OK f' => if check_valid_node (fn_code f') (RTL.fn_entrypoint f0) then if forall_ptree (check_code (RTL.fn_code f0) (fn_code f')) (fn_code f') then OK {| fn_sig := RTL.fn_sig f0; fn_params := RTL.fn_params f0; fn_stacksize := RTL.fn_stacksize f0; fn_code := fn_code f'; fn_entrypoint := RTL.fn_entrypoint f0 |} else Error (msg "check_present_blocks failed") else Error (msg "entrypoint does not exists") | Error msg => Error msg end = OK f1

funsig (Internal f1) = RTL.funsig (Internal f0)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.fundef
tf: fundef
e: external_function
Heqf0: f = External e
H: OK (External e) = OK tf
funsig tf = RTL.funsig (External e)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f0: RTL.function
f1, f: function
Heqr0: partition f0 = OK f
Heqb: check_valid_node (fn_code f) (RTL.fn_entrypoint f0) = true
Heqb0: forall_ptree (check_code (RTL.fn_code f0) (fn_code f)) (fn_code f) = true
Heqr: OK {| fn_sig := RTL.fn_sig f0; fn_params := RTL.fn_params f0; fn_stacksize := RTL.fn_stacksize f0; fn_code := fn_code f; fn_entrypoint := RTL.fn_entrypoint f0 |} = OK f1

funsig (Internal f1) = RTL.funsig (Internal f0)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.fundef
tf: fundef
e: external_function
Heqf0: f = External e
H: OK (External e) = OK tf
funsig tf = RTL.funsig (External e)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f0: RTL.function
f: function
Heqr0: partition f0 = OK f
Heqb: check_valid_node (fn_code f) (RTL.fn_entrypoint f0) = true
Heqb0: forall_ptree (check_code (RTL.fn_code f0) (fn_code f)) (fn_code f) = true

funsig (Internal {| fn_sig := RTL.fn_sig f0; fn_params := RTL.fn_params f0; fn_stacksize := RTL.fn_stacksize f0; fn_code := fn_code f; fn_entrypoint := RTL.fn_entrypoint f0 |}) = RTL.funsig (Internal f0)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.fundef
tf: fundef
e: external_function
Heqf0: f = External e
H: OK (External e) = OK tf
funsig tf = RTL.funsig (External e)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.fundef
tf: fundef
e: external_function
Heqf0: f = External e
H: OK (External e) = OK tf

funsig tf = RTL.funsig (External e)
inv H; auto. Qed. Definition measure (b: option SeqBB.t): nat := match b with | None => 0 | Some b' => 1 + length b' end.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (a : code) (b : node), check_valid_node a b = true -> valid_succ a b
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (a : code) (b : node), check_valid_node a b = true -> valid_succ a b
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
a: code
b: node
H: check_valid_node a b = true

valid_succ a b
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
a: code
b: node
H: match a ! b with | Some _ => true | None => false end = true

exists b0 : SeqBB.t, a ! b = Some b0
destruct_match; try discriminate; eauto. Qed.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (f : RTL.function) (tf : function), transl_function f = OK tf -> valid_succ (fn_code tf) (fn_entrypoint tf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (f : RTL.function) (tf : function), transl_function f = OK tf -> valid_succ (fn_code tf) (fn_entrypoint tf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.function
tf: function
H: match partition f with | OK f' => if check_valid_node (fn_code f') (RTL.fn_entrypoint f) then if forall_ptree (check_code (RTL.fn_code f) (fn_code f')) (fn_code f') then OK {| fn_sig := RTL.fn_sig f; fn_params := RTL.fn_params f; fn_stacksize := RTL.fn_stacksize f; fn_code := fn_code f'; fn_entrypoint := RTL.fn_entrypoint f |} else Error (msg "check_present_blocks failed") else Error (msg "entrypoint does not exists") | Error msg => Error msg end = OK tf

valid_succ (fn_code tf) (fn_entrypoint tf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.function
tf, f0: function
Heqr: partition f = OK f0
Heqb: check_valid_node (fn_code f0) (RTL.fn_entrypoint f) = true
Heqb0: forall_ptree (check_code (RTL.fn_code f) (fn_code f0)) (fn_code f0) = true
H: OK {| fn_sig := RTL.fn_sig f; fn_params := RTL.fn_params f; fn_stacksize := RTL.fn_stacksize f; fn_code := fn_code f0; fn_entrypoint := RTL.fn_entrypoint f |} = OK tf

valid_succ (fn_code tf) (fn_entrypoint tf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.function
f0: function
Heqr: partition f = OK f0
Heqb: check_valid_node (fn_code f0) (RTL.fn_entrypoint f) = true
Heqb0: forall_ptree (check_code (RTL.fn_code f) (fn_code f0)) (fn_code f0) = true

valid_succ (fn_code {| fn_sig := RTL.fn_sig f; fn_params := RTL.fn_params f; fn_stacksize := RTL.fn_stacksize f; fn_code := fn_code f0; fn_entrypoint := RTL.fn_entrypoint f |}) (fn_entrypoint {| fn_sig := RTL.fn_sig f; fn_params := RTL.fn_params f; fn_stacksize := RTL.fn_stacksize f; fn_code := fn_code f0; fn_entrypoint := RTL.fn_entrypoint f |})
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.function
f0: function
Heqr: partition f = OK f0
Heqb: check_valid_node (fn_code f0) (RTL.fn_entrypoint f) = true
Heqb0: forall_ptree (check_code (RTL.fn_code f) (fn_code f0)) (fn_code f0) = true

valid_succ (fn_code f0) (RTL.fn_entrypoint f)
eauto using check_valid_node_correct. Qed.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (A : Type) (a : forall a b : A, {a = b} + {a <> b}) (b c : A), ceq a b c = true -> b = c
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (A : Type) (a : forall a b : A, {a = b} + {a <> b}) (b c : A), ceq a b c = true -> b = c
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
A: Type
a: forall a b : A, {a = b} + {a <> b}
b, c: A
H: ceq a b c = true

b = c
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
A: Type
a: forall a b : A, {a = b} + {a <> b}
b, c: A
H: (if a b c then true else false) = true

b = c
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
A: Type
a: forall a b : A, {a = b} + {a <> b}
b, c: A
e: b = c
Heqs: a b c = left e
H: true = true

b = c
auto. Qed. Ltac unfold_ands := repeat match goal with | H: _ && _ = true |- _ => apply andb_prop in H | H: _ /\ _ |- _ => inv H | H: ceq _ _ _ = true |- _ => apply ceq_eq in H; subst end.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (c : RTL.code) (tc : code) (b : SeqBB.t) (pc : node), check_code c tc pc b = true -> match_block c tc pc b
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (c : RTL.code) (tc : code) (b : SeqBB.t) (pc : node), check_code c tc pc b = true -> match_block c tc pc b
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
pc: node
H: match c ! pc with | Some (RTL.Inop _) | _ => false end = true

match_block c tc pc nil
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
forall pc : node, check_code c tc pc (a :: b) = true -> match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
pc: node
H: match c ! pc with | Some (RTL.Inop _) | _ => false end = true

match_block c tc pc nil
repeat (destruct_match; try discriminate).
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b

forall pc : node, check_code c tc pc (a :: b) = true -> match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
H: check_code c tc pc (a :: b) = true

match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
H: match c ! pc with | Some (RTL.Inop pc'') => match a with | RBnop => match b with | RBop _ _ _ _ :: _ :: _ | RBload _ _ _ _ _ :: _ :: _ | RBstore _ _ _ _ _ :: _ :: _ | RBexit None (RBcall _ _ _ _ _) :: _ :: _ | RBexit None (RBjumptable _ _) :: _ :: _ => (fix check_code (c : RTL.code) (tc : code) (pc : node) (b : SeqBB.t) {struct b} : bool := match c ! pc with | Some (RTL.Inop pc''0) => match b with | RBnop :: (RBnop :: _ :: _) as b' | RBnop :: (RBop _ _ _ _ :: _ :: _) as b' | RBnop :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBnop :: (RBexit (Some _) _ :: _ :: _) as b' | RBnop :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBreturn _) :: _ :: _) as b' => check_code c tc pc''0 b' | RBnop :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 | RBnop :: (RBexit None (RBgoto pc') :: _ :: _) as b' => check_code c tc pc''0 b' | _ => false end | Some (RTL.Iop op' args' dst' pc''0) => match b with | RBop None op args dst :: nil | RBop None op args dst :: RBnop :: nil | RBop None op args dst :: RBop _ _ _ _ :: nil | RBop None op args dst :: RBload _ _ _ _ _ :: nil | RBop None op args dst :: RBstore _ _ _ _ _ :: nil | RBop None op args dst :: RBsetpred _ _ _ _ :: nil | RBop None op args dst :: RBexit (Some _) _ :: nil | RBop None op args dst :: RBexit None (RBcall _ _ _ _ _) :: nil | RBop None op args dst :: RBexit None (RBtailcall _ _ _) :: nil | RBop None op args dst :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBop None op args dst :: RBexit None (RBcond _ _ _ _) :: nil | RBop None op args dst :: RBexit None (RBjumptable _ _) :: nil | RBop None op args dst :: RBexit None (RBreturn _) :: nil => false | RBop None op args dst :: (RBnop :: _ :: _) as b' | RBop None op args dst :: (RBop _ _ _ _ :: _ :: _) as b' | RBop None op args dst :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBop None op args dst :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBop None op args dst :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBop None op args dst :: (RBexit (Some _) _ :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args args' && ceq peq dst dst' && check_code c tc pc''0 b' | RBop None op args dst :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq operation_eq op op' && ceq list_pos_eq args args' && ceq peq dst dst' | RBop None op args dst :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args args' && ceq peq dst dst' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Iload chunk' addr' args' dst' pc''0) => match b with | RBload None chunk addr args dst :: nil | RBload None chunk addr args dst :: RBnop :: nil | RBload None chunk addr args dst :: RBop _ _ _ _ :: nil | RBload None chunk addr args dst :: RBload _ _ _ _ _ :: nil | RBload None chunk addr args dst :: RBstore _ _ _ _ _ :: nil | RBload None chunk addr args dst :: RBsetpred _ _ _ _ :: nil | RBload None chunk addr args dst :: RBexit (Some _) _ :: nil | RBload None chunk addr args dst :: RBexit None (RBcall _ _ _ _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBtailcall _ _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBcond _ _ _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBjumptable _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBreturn _) :: nil => false | RBload None chunk addr args dst :: (RBnop :: _ :: _) as b' | RBload None chunk addr args dst :: (RBop _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit (Some _) _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq dst dst' && check_code c tc pc''0 b' | RBload None chunk addr args dst :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq dst dst' | RBload None chunk addr args dst :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq dst dst' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Istore chunk' addr' args' src' pc''0) => match b with | RBstore None chunk addr args src :: nil | RBstore None chunk addr args src :: RBnop :: nil | RBstore None chunk addr args src :: RBop _ _ _ _ :: nil | RBstore None chunk addr args src :: RBload _ _ _ _ _ :: nil | RBstore None chunk addr args src :: RBstore _ _ _ _ _ :: nil | RBstore None chunk addr args src :: RBsetpred _ _ _ _ :: nil | RBstore None chunk addr args src :: RBexit (Some _) _ :: nil | RBstore None chunk addr args src :: RBexit None (RBcall _ _ _ _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBtailcall _ _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBcond _ _ _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBjumptable _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBreturn _) :: nil => false | RBstore None chunk addr args src :: (RBnop :: _ :: _) as b' | RBstore None chunk addr args src :: (RBop _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit (Some _) _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq src src' && check_code c tc pc''0 b' | RBstore None chunk addr args src :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq src src' | RBstore None chunk addr args src :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq src src' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Icall sig' (inl r') args' dst' pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl r0) args dst pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq signature_eq sig sig' && ceq peq r0 r' && ceq list_pos_eq args args' && ceq peq dst dst' | RBnop :: RBexit None (RBcall sig (inl r0) args dst pc') :: _ :: _ => false | RBnop :: RBexit None (RBcall sig (inr _) _ _ _) :: _ => false | _ => false end | Some (RTL.Icall sig' (inr i') args' dst' pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl _) _ _ _) :: _ => false | RBnop :: RBexit None (RBcall sig (inr i) args dst pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq signature_eq sig sig' && ceq peq i i' && ceq list_pos_eq args args' && ceq peq dst dst' | RBnop :: RBexit None (RBcall sig (inr i) args dst pc') :: _ :: _ => false | _ => false end | Some (RTL.Itailcall sig (inl r0) args) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl r') args') :: nil => ceq signature_eq sig sig' && ceq peq r0 r' && ceq list_pos_eq args args' | RBnop :: RBexit None (RBtailcall sig' (inl r') args') :: _ :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr _) _) :: _ => false | _ => false end | Some (RTL.Itailcall sig (inr r0) args) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl _) _) :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr r') args') :: nil => ceq signature_eq sig sig' && ceq peq r0 r' && ceq list_pos_eq args args' | RBnop :: RBexit None (RBtailcall sig' (inr r') args') :: _ :: _ => false | _ => false end | Some (RTL.Icond cond args n1 n2) => match b with | RBnop :: RBexit None (RBcond cond' args' n1' n2') :: nil => check_valid_node tc n1 && check_valid_node tc n2 && ceq condition_eq cond cond' && ceq list_pos_eq args args' && ceq peq n1 n1' && ceq peq n2 n2' | RBnop :: RBexit None (RBcond cond' args' n1' n2') :: _ :: _ => false | _ => false end | Some (RTL.Ijumptable r0 ns) => match b with | RBnop :: RBexit None (RBjumptable r' ns') :: nil => ceq peq r0 r' && ceq list_pos_eq ns ns' && forallb (check_valid_node tc) ns | RBnop :: RBexit None (RBjumptable r' ns') :: _ :: _ => false | _ => false end | Some (RTL.Ireturn (Some r0)) => match b with | RBnop :: RBexit None (RBreturn (Some r')) :: nil => ceq peq r0 r' | RBnop :: RBexit None (RBreturn (Some r')) :: _ :: _ => false | _ => false end | Some (RTL.Ireturn None) => match b with | RBnop :: RBexit None (RBreturn None) :: nil => true | _ => false end | _ => false end) c tc pc'' b | RBnop :: _ :: _ | RBsetpred _ _ _ _ :: _ :: _ | RBexit (Some _) _ :: _ :: _ | RBexit None (RBtailcall _ _ _) :: _ :: _ | RBexit None (RBbuiltin _ _ _ _) :: _ :: _ | RBexit None (RBcond _ _ _ _) :: _ :: _ | RBexit None (RBreturn _) :: _ :: _ => (fix check_code (c : RTL.code) (tc : code) (pc : node) (b : SeqBB.t) {struct b} : bool := match c ! pc with | Some (RTL.Inop pc''0) => match b with | RBnop :: (RBnop :: _ :: _) as b' | RBnop :: (RBop _ _ _ _ :: _ :: _) as b' | RBnop :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBnop :: (RBexit (Some _) _ :: _ :: _) as b' | RBnop :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBreturn _) :: _ :: _) as b' => check_code c tc pc''0 b' | RBnop :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 | RBnop :: (RBexit None (RBgoto pc') :: _ :: _) as b' => check_code c tc pc''0 b' | _ => false end | Some (RTL.Iop op' args' dst' pc''0) => match b with | RBop None op args dst :: nil | RBop None op args dst :: RBnop :: nil | RBop None op args dst :: RBop _ _ _ _ :: nil | RBop None op args dst :: RBload _ _ _ _ _ :: nil | RBop None op args dst :: RBstore _ _ _ _ _ :: nil | RBop None op args dst :: RBsetpred _ _ _ _ :: nil | RBop None op args dst :: RBexit (Some _) _ :: nil | RBop None op args dst :: RBexit None (RBcall _ _ _ _ _) :: nil | RBop None op args dst :: RBexit None (RBtailcall _ _ _) :: nil | RBop None op args dst :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBop None op args dst :: RBexit None (RBcond _ _ _ _) :: nil | RBop None op args dst :: RBexit None (RBjumptable _ _) :: nil | RBop None op args dst :: RBexit None (RBreturn _) :: nil => false | RBop None op args dst :: (RBnop :: _ :: _) as b' | RBop None op args dst :: (RBop _ _ _ _ :: _ :: _) as b' | RBop None op args dst :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBop None op args dst :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBop None op args dst :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBop None op args dst :: (RBexit (Some _) _ :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args args' && ceq peq dst dst' && check_code c tc pc''0 b' | RBop None op args dst :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq operation_eq op op' && ceq list_pos_eq args args' && ceq peq dst dst' | RBop None op args dst :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args args' && ceq peq dst dst' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Iload chunk' addr' args' dst' pc''0) => match b with | RBload None chunk addr args dst :: nil | RBload None chunk addr args dst :: RBnop :: nil | RBload None chunk addr args dst :: RBop _ _ _ _ :: nil | RBload None chunk addr args dst :: RBload _ _ _ _ _ :: nil | RBload None chunk addr args dst :: RBstore _ _ _ _ _ :: nil | RBload None chunk addr args dst :: RBsetpred _ _ _ _ :: nil | RBload None chunk addr args dst :: RBexit (Some _) _ :: nil | RBload None chunk addr args dst :: RBexit None (RBcall _ _ _ _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBtailcall _ _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBcond _ _ _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBjumptable _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBreturn _) :: nil => false | RBload None chunk addr args dst :: (RBnop :: _ :: _) as b' | RBload None chunk addr args dst :: (RBop _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit (Some _) _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq dst dst' && check_code c tc pc''0 b' | RBload None chunk addr args dst :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq dst dst' | RBload None chunk addr args dst :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq dst dst' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Istore chunk' addr' args' src' pc''0) => match b with | RBstore None chunk addr args src :: nil | RBstore None chunk addr args src :: RBnop :: nil | RBstore None chunk addr args src :: RBop _ _ _ _ :: nil | RBstore None chunk addr args src :: RBload _ _ _ _ _ :: nil | RBstore None chunk addr args src :: RBstore _ _ _ _ _ :: nil | RBstore None chunk addr args src :: RBsetpred _ _ _ _ :: nil | RBstore None chunk addr args src :: RBexit (Some _) _ :: nil | RBstore None chunk addr args src :: RBexit None (RBcall _ _ _ _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBtailcall _ _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBcond _ _ _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBjumptable _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBreturn _) :: nil => false | RBstore None chunk addr args src :: (RBnop :: _ :: _) as b' | RBstore None chunk addr args src :: (RBop _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit (Some _) _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq src src' && check_code c tc pc''0 b' | RBstore None chunk addr args src :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq src src' | RBstore None chunk addr args src :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq src src' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Icall sig' (inl r') args' dst' pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl r) args dst pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args args' && ceq peq dst dst' | RBnop :: RBexit None (RBcall sig (inl r) args dst pc') :: _ :: _ => false | RBnop :: RBexit None (RBcall sig (inr _) _ _ _) :: _ => false | _ => false end | Some (RTL.Icall sig' (inr i') args' dst' pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl _) _ _ _) :: _ => false | RBnop :: RBexit None (RBcall sig (inr i) args dst pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq signature_eq sig sig' && ceq peq i i' && ceq list_pos_eq args args' && ceq peq dst dst' | RBnop :: RBexit None (RBcall sig (inr i) args dst pc') :: _ :: _ => false | _ => false end | Some (RTL.Itailcall sig (inl r) args) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl r') args') :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args args' | RBnop :: RBexit None (RBtailcall sig' (inl r') args') :: _ :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr _) _) :: _ => false | _ => false end | Some (RTL.Itailcall sig (inr r) args) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl _) _) :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr r') args') :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args args' | RBnop :: RBexit None (RBtailcall sig' (inr r') args') :: _ :: _ => false | _ => false end | Some (RTL.Icond cond args n1 n2) => match b with | RBnop :: RBexit None (RBcond cond' args' n1' n2') :: nil => check_valid_node tc n1 && check_valid_node tc n2 && ceq condition_eq cond cond' && ceq list_pos_eq args args' && ceq peq n1 n1' && ceq peq n2 n2' | RBnop :: RBexit None (RBcond cond' args' n1' n2') :: _ :: _ => false | _ => false end | Some (RTL.Ijumptable r ns) => match b with | RBnop :: RBexit None (RBjumptable r' ns') :: nil => ceq peq r r' && ceq list_pos_eq ns ns' && forallb (check_valid_node tc) ns | RBnop :: RBexit None (RBjumptable r' ns') :: _ :: _ => false | _ => false end | Some (RTL.Ireturn (Some r)) => match b with | RBnop :: RBexit None (RBreturn (Some r')) :: nil => ceq peq r r' | RBnop :: RBexit None (RBreturn (Some r')) :: _ :: _ => false | _ => false end | Some (RTL.Ireturn None) => match b with | RBnop :: RBexit None (RBreturn None) :: nil => true | _ => false end | _ => false end) c tc pc'' b | RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc'' | RBexit None (RBgoto pc') :: _ :: _ => (fix check_code (c : RTL.code) (tc : code) (pc : node) (b : SeqBB.t) {struct b} : bool := match c ! pc with | Some (RTL.Inop pc''0) => match b with | RBnop :: (RBnop :: _ :: _) as b' | RBnop :: (RBop _ _ _ _ :: _ :: _) as b' | RBnop :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBnop :: (RBexit (Some _) _ :: _ :: _) as b' | RBnop :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBreturn _) :: _ :: _) as b' => check_code c tc pc''0 b' | RBnop :: RBexit None (RBgoto pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 | RBnop :: (RBexit None (RBgoto pc'0) :: _ :: _) as b' => check_code c tc pc''0 b' | _ => false end | Some (RTL.Iop op' args' dst' pc''0) => match b with | RBop None op args dst :: nil | RBop None op args dst :: RBnop :: nil | RBop None op args dst :: RBop _ _ _ _ :: nil | RBop None op args dst :: RBload _ _ _ _ _ :: nil | RBop None op args dst :: RBstore _ _ _ _ _ :: nil | RBop None op args dst :: RBsetpred _ _ _ _ :: nil | RBop None op args dst :: RBexit (Some _) _ :: nil | RBop None op args dst :: RBexit None (RBcall _ _ _ _ _) :: nil | RBop None op args dst :: RBexit None (RBtailcall _ _ _) :: nil | RBop None op args dst :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBop None op args dst :: RBexit None (RBcond _ _ _ _) :: nil | RBop None op args dst :: RBexit None (RBjumptable _ _) :: nil | RBop None op args dst :: RBexit None (RBreturn _) :: nil => false | RBop None op args dst :: (RBnop :: _ :: _) as b' | RBop None op args dst :: (RBop _ _ _ _ :: _ :: _) as b' | RBop None op args dst :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBop None op args dst :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBop None op args dst :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBop None op args dst :: (RBexit (Some _) _ :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args args' && ceq peq dst dst' && check_code c tc pc''0 b' | RBop None op args dst :: RBexit None (RBgoto pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq operation_eq op op' && ceq list_pos_eq args args' && ceq peq dst dst' | RBop None op args dst :: (RBexit None (RBgoto pc'0) :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args args' && ceq peq dst dst' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Iload chunk' addr' args' dst' pc''0) => match b with | RBload None chunk addr args dst :: nil | RBload None chunk addr args dst :: RBnop :: nil | RBload None chunk addr args dst :: RBop _ _ _ _ :: nil | RBload None chunk addr args dst :: RBload _ _ _ _ _ :: nil | RBload None chunk addr args dst :: RBstore _ _ _ _ _ :: nil | RBload None chunk addr args dst :: RBsetpred _ _ _ _ :: nil | RBload None chunk addr args dst :: RBexit (Some _) _ :: nil | RBload None chunk addr args dst :: RBexit None (RBcall _ _ _ _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBtailcall _ _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBcond _ _ _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBjumptable _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBreturn _) :: nil => false | RBload None chunk addr args dst :: (RBnop :: _ :: _) as b' | RBload None chunk addr args dst :: (RBop _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit (Some _) _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq dst dst' && check_code c tc pc''0 b' | RBload None chunk addr args dst :: RBexit None (RBgoto pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq dst dst' | RBload None chunk addr args dst :: (RBexit None (RBgoto pc'0) :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq dst dst' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Istore chunk' addr' args' src' pc''0) => match b with | RBstore None chunk addr args src :: nil | RBstore None chunk addr args src :: RBnop :: nil | RBstore None chunk addr args src :: RBop _ _ _ _ :: nil | RBstore None chunk addr args src :: RBload _ _ _ _ _ :: nil | RBstore None chunk addr args src :: RBstore _ _ _ _ _ :: nil | RBstore None chunk addr args src :: RBsetpred _ _ _ _ :: nil | RBstore None chunk addr args src :: RBexit (Some _) _ :: nil | RBstore None chunk addr args src :: RBexit None (RBcall _ _ _ _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBtailcall _ _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBcond _ _ _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBjumptable _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBreturn _) :: nil => false | RBstore None chunk addr args src :: (RBnop :: _ :: _) as b' | RBstore None chunk addr args src :: (RBop _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit (Some _) _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq src src' && check_code c tc pc''0 b' | RBstore None chunk addr args src :: RBexit None (RBgoto pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq src src' | RBstore None chunk addr args src :: (RBexit None (RBgoto pc'0) :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq src src' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Icall sig' (inl r') args' dst' pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl r) args dst pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args args' && ceq peq dst dst' | RBnop :: RBexit None (RBcall sig (inl r) args dst pc'0) :: _ :: _ => false | RBnop :: RBexit None (RBcall sig (inr _) _ _ _) :: _ => false | _ => false end | Some (RTL.Icall sig' (inr i') args' dst' pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl _) _ _ _) :: _ => false | RBnop :: RBexit None (RBcall sig (inr i) args dst pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq signature_eq sig sig' && ceq peq i i' && ceq list_pos_eq args args' && ceq peq dst dst' | RBnop :: RBexit None (RBcall sig (inr i) args dst pc'0) :: _ :: _ => false | _ => false end | Some (RTL.Itailcall sig (inl r) args) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl r') args') :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args args' | RBnop :: RBexit None (RBtailcall sig' (inl r') args') :: _ :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr _) _) :: _ => false | _ => false end | Some (RTL.Itailcall sig (inr r) args) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl _) _) :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr r') args') :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args args' | RBnop :: RBexit None (RBtailcall sig' (inr r') args') :: _ :: _ => false | _ => false end | Some (RTL.Icond cond args n1 n2) => match b with | RBnop :: RBexit None (RBcond cond' args' n1' n2') :: nil => check_valid_node tc n1 && check_valid_node tc n2 && ceq condition_eq cond cond' && ceq list_pos_eq args args' && ceq peq n1 n1' && ceq peq n2 n2' | RBnop :: RBexit None (RBcond cond' args' n1' n2') :: _ :: _ => false | _ => false end | Some (RTL.Ijumptable r ns) => match b with | RBnop :: RBexit None (RBjumptable r' ns') :: nil => ceq peq r r' && ceq list_pos_eq ns ns' && forallb (check_valid_node tc) ns | RBnop :: RBexit None (RBjumptable r' ns') :: _ :: _ => false | _ => false end | Some (RTL.Ireturn (Some r)) => match b with | RBnop :: RBexit None (RBreturn (Some r')) :: nil => ceq peq r r' | RBnop :: RBexit None (RBreturn (Some r')) :: _ :: _ => false | _ => false end | Some (RTL.Ireturn None) => match b with | RBnop :: RBexit None (RBreturn None) :: nil => true | _ => false end | _ => false end) c tc pc'' b | _ => false end | _ => false end | Some (RTL.Iop op' args' dst' pc'') => match a with | RBop None op args dst => match b with | RBop _ _ _ _ :: _ :: _ | RBload _ _ _ _ _ :: _ :: _ | RBstore _ _ _ _ _ :: _ :: _ | RBexit None (RBcall _ _ _ _ _) :: _ :: _ | RBexit None (RBjumptable _ _) :: _ :: _ => ceq operation_eq op op' && ceq list_pos_eq args args' && ceq peq dst dst' && (fix check_code (c : RTL.code) (tc : code) (pc : node) (b : SeqBB.t) {struct b} : bool := match c ! pc with | Some (RTL.Inop pc''0) => match b with | RBnop :: (RBnop :: _ :: _) as b' | RBnop :: (RBop _ _ _ _ :: _ :: _) as b' | RBnop :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBnop :: (RBexit (Some _) _ :: _ :: _) as b' | RBnop :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBreturn _) :: _ :: _) as b' => check_code c tc pc''0 b' | RBnop :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 | RBnop :: (RBexit None (RBgoto pc') :: _ :: _) as b' => check_code c tc pc''0 b' | _ => false end | Some (RTL.Iop op'0 args'0 dst'0 pc''0) => match b with | RBop None op0 args0 dst0 :: nil | RBop None op0 args0 dst0 :: RBnop :: nil | RBop None op0 args0 dst0 :: RBop _ _ _ _ :: nil | RBop None op0 args0 dst0 :: RBload _ _ _ _ _ :: nil | RBop None op0 args0 dst0 :: RBstore _ _ _ _ _ :: nil | RBop None op0 args0 dst0 :: RBsetpred _ _ _ _ :: nil | RBop None op0 args0 dst0 :: RBexit (Some _) _ :: nil | RBop None op0 args0 dst0 :: RBexit None (RBcall _ _ _ _ _) :: nil | RBop None op0 args0 dst0 :: RBexit None (RBtailcall _ _ _) :: nil | RBop None op0 args0 dst0 :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBop None op0 args0 dst0 :: RBexit None (RBcond _ _ _ _) :: nil | RBop None op0 args0 dst0 :: RBexit None (RBjumptable _ _) :: nil | RBop None op0 args0 dst0 :: RBexit None (RBreturn _) :: nil => false | RBop None op0 args0 dst0 :: (RBnop :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBop _ _ _ _ :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit (Some _) _ :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq operation_eq op0 op'0 && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | RBop None op0 args0 dst0 :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq operation_eq op0 op'0 && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBop None op0 args0 dst0 :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq operation_eq op0 op'0 && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | _ => false end | Some (RTL.Iload chunk' addr' args'0 dst'0 pc''0) => match b with | RBload None chunk addr args0 dst0 :: nil | RBload None chunk addr args0 dst0 :: RBnop :: nil | RBload None chunk addr args0 dst0 :: RBop _ _ _ _ :: nil | RBload None chunk addr args0 dst0 :: RBload _ _ _ _ _ :: nil | RBload None chunk addr args0 dst0 :: RBstore _ _ _ _ _ :: nil | RBload None chunk addr args0 dst0 :: RBsetpred _ _ _ _ :: nil | RBload None chunk addr args0 dst0 :: RBexit (Some _) _ :: nil | RBload None chunk addr args0 dst0 :: RBexit None (RBcall _ _ _ _ _) :: nil | RBload None chunk addr args0 dst0 :: RBexit None (RBtailcall _ _ _) :: nil | RBload None chunk addr args0 dst0 :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBload None chunk addr args0 dst0 :: RBexit None (RBcond _ _ _ _) :: nil | RBload None chunk addr args0 dst0 :: RBexit None (RBjumptable _ _) :: nil | RBload None chunk addr args0 dst0 :: RBexit None (RBreturn _) :: nil => false | RBload None chunk addr args0 dst0 :: (RBnop :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBop _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit (Some _) _ :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | RBload None chunk addr args0 dst0 :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBload None chunk addr args0 dst0 :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | _ => false end | Some (RTL.Istore chunk' addr' args'0 src' pc''0) => match b with | RBstore None chunk addr args0 src :: nil | RBstore None chunk addr args0 src :: RBnop :: nil | RBstore None chunk addr args0 src :: RBop _ _ _ _ :: nil | RBstore None chunk addr args0 src :: RBload _ _ _ _ _ :: nil | RBstore None chunk addr args0 src :: RBstore _ _ _ _ _ :: nil | RBstore None chunk addr args0 src :: RBsetpred _ _ _ _ :: nil | RBstore None chunk addr args0 src :: RBexit (Some _) _ :: nil | RBstore None chunk addr args0 src :: RBexit None (RBcall _ _ _ _ _) :: nil | RBstore None chunk addr args0 src :: RBexit None (RBtailcall _ _ _) :: nil | RBstore None chunk addr args0 src :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBstore None chunk addr args0 src :: RBexit None (RBcond _ _ _ _) :: nil | RBstore None chunk addr args0 src :: RBexit None (RBjumptable _ _) :: nil | RBstore None chunk addr args0 src :: RBexit None (RBreturn _) :: nil => false | RBstore None chunk addr args0 src :: (RBnop :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBop _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit (Some _) _ :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args0 args'0 && ceq peq src src' && check_code c tc pc''0 b' | RBstore None chunk addr args0 src :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args0 args'0 && ceq peq src src' | RBstore None chunk addr args0 src :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args0 args'0 && ceq peq src src' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Icall sig' (inl r') args'0 dst'0 pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl r0) args0 dst0 pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq signature_eq sig sig' && ceq peq r0 r' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBnop :: RBexit None (RBcall sig (inl r0) args0 dst0 pc') :: _ :: _ => false | RBnop :: RBexit None (RBcall sig (inr _) _ _ _) :: _ => false | _ => false end | Some (RTL.Icall sig' (inr i') args'0 dst'0 pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl _) _ _ _) :: _ => false | RBnop :: RBexit None (RBcall sig (inr i) args0 dst0 pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq signature_eq sig sig' && ceq peq i i' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBnop :: RBexit None (RBcall sig (inr i) args0 dst0 pc') :: _ :: _ => false | _ => false end | Some (RTL.Itailcall sig (inl r0) args0) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl r') args'0) :: nil => ceq signature_eq sig sig' && ceq peq r0 r' && ceq list_pos_eq args0 args'0 | RBnop :: RBexit None (RBtailcall sig' (inl r') args'0) :: _ :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr _) _) :: _ => false | _ => false end | Some (RTL.Itailcall sig (inr r0) args0) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl _) _) :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr r') args'0) :: nil => ceq signature_eq sig sig' && ceq peq r0 r' && ceq list_pos_eq args0 args'0 | RBnop :: RBexit None (RBtailcall sig' (inr r') args'0) :: _ :: _ => false | _ => false end | Some (RTL.Icond cond args0 n1 n2) => match b with | RBnop :: RBexit None (RBcond cond' args'0 n1' n2') :: nil => check_valid_node tc n1 && check_valid_node tc n2 && ceq condition_eq cond cond' && ceq list_pos_eq args0 args'0 && ceq peq n1 n1' && ceq peq n2 n2' | RBnop :: RBexit None (RBcond cond' args'0 n1' n2') :: _ :: _ => false | _ => false end | Some (RTL.Ijumptable r0 ns) => match b with | RBnop :: RBexit None (RBjumptable r' ns') :: nil => ceq peq r0 r' && ceq list_pos_eq ns ns' && forallb (check_valid_node tc) ns | RBnop :: RBexit None (RBjumptable r' ns') :: _ :: _ => false | _ => false end | Some (RTL.Ireturn (Some r0)) => match b with | RBnop :: RBexit None (RBreturn (Some r')) :: nil => ceq peq r0 r' | RBnop :: RBexit None (RBreturn (Some r')) :: _ :: _ => false | _ => false end | Some (RTL.Ireturn None) => match b with | RBnop :: RBexit None (RBreturn None) :: nil => true | _ => false end | _ => false end) c tc pc'' b | RBnop :: _ :: _ | RBsetpred _ _ _ _ :: _ :: _ | RBexit (Some _) _ :: _ :: _ | RBexit None (RBtailcall _ _ _) :: _ :: _ | RBexit None (RBbuiltin _ _ _ _) :: _ :: _ | RBexit None (RBcond _ _ _ _) :: _ :: _ | RBexit None (RBreturn _) :: _ :: _ => ceq operation_eq op op' && ceq list_pos_eq args args' && ceq peq dst dst' && (fix check_code (c : RTL.code) (tc : code) (pc : node) (b : SeqBB.t) {struct b} : bool := match c ! pc with | Some (RTL.Inop pc''0) => match b with | RBnop :: (RBnop :: _ :: _) as b' | RBnop :: (RBop _ _ _ _ :: _ :: _) as b' | RBnop :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBnop :: (RBexit (Some _) _ :: _ :: _) as b' | RBnop :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBreturn _) :: _ :: _) as b' => check_code c tc pc''0 b' | RBnop :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 | RBnop :: (RBexit None (RBgoto pc') :: _ :: _) as b' => check_code c tc pc''0 b' | _ => false end | Some (RTL.Iop op'0 args'0 dst'0 pc''0) => match b with | RBop None op0 args0 dst0 :: nil | RBop None op0 args0 dst0 :: RBnop :: nil | RBop None op0 args0 dst0 :: RBop _ _ _ _ :: nil | RBop None op0 args0 dst0 :: RBload _ _ _ _ _ :: nil | RBop None op0 args0 dst0 :: RBstore _ _ _ _ _ :: nil | RBop None op0 args0 dst0 :: RBsetpred _ _ _ _ :: nil | RBop None op0 args0 dst0 :: RBexit (Some _) _ :: nil | RBop None op0 args0 dst0 :: RBexit None (RBcall _ _ _ _ _) :: nil | RBop None op0 args0 dst0 :: RBexit None (RBtailcall _ _ _) :: nil | RBop None op0 args0 dst0 :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBop None op0 args0 dst0 :: RBexit None (RBcond _ _ _ _) :: nil | RBop None op0 args0 dst0 :: RBexit None (RBjumptable _ _) :: nil | RBop None op0 args0 dst0 :: RBexit None (RBreturn _) :: nil => false | RBop None op0 args0 dst0 :: (RBnop :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBop _ _ _ _ :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit (Some _) _ :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq operation_eq op0 op'0 && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | RBop None op0 args0 dst0 :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq operation_eq op0 op'0 && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBop None op0 args0 dst0 :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq operation_eq op0 op'0 && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | _ => false end | Some (RTL.Iload chunk' addr' args'0 dst'0 pc''0) => match b with | RBload None chunk addr args0 dst0 :: nil | RBload None chunk addr args0 dst0 :: RBnop :: nil | RBload None chunk addr args0 dst0 :: RBop _ _ _ _ :: nil | RBload None chunk addr args0 dst0 :: RBload _ _ _ _ _ :: nil | RBload None chunk addr args0 dst0 :: RBstore _ _ _ _ _ :: nil | RBload None chunk addr args0 dst0 :: RBsetpred _ _ _ _ :: nil | RBload None chunk addr args0 dst0 :: RBexit (Some _) _ :: nil | RBload None chunk addr args0 dst0 :: RBexit None (RBcall _ _ _ _ _) :: nil | RBload None chunk addr args0 dst0 :: RBexit None (RBtailcall _ _ _) :: nil | RBload None chunk addr args0 dst0 :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBload None chunk addr args0 dst0 :: RBexit None (RBcond _ _ _ _) :: nil | RBload None chunk addr args0 dst0 :: RBexit None (RBjumptable _ _) :: nil | RBload None chunk addr args0 dst0 :: RBexit None (RBreturn _) :: nil => false | RBload None chunk addr args0 dst0 :: (RBnop :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBop _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit (Some _) _ :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | RBload None chunk addr args0 dst0 :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBload None chunk addr args0 dst0 :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | _ => false end | Some (RTL.Istore chunk' addr' args'0 src' pc''0) => match b with | RBstore None chunk addr args0 src :: nil | RBstore None chunk addr args0 src :: RBnop :: nil | RBstore None chunk addr args0 src :: RBop _ _ _ _ :: nil | RBstore None chunk addr args0 src :: RBload _ _ _ _ _ :: nil | RBstore None chunk addr args0 src :: RBstore _ _ _ _ _ :: nil | RBstore None chunk addr args0 src :: RBsetpred _ _ _ _ :: nil | RBstore None chunk addr args0 src :: RBexit (Some _) _ :: nil | RBstore None chunk addr args0 src :: RBexit None (RBcall _ _ _ _ _) :: nil | RBstore None chunk addr args0 src :: RBexit None (RBtailcall _ _ _) :: nil | RBstore None chunk addr args0 src :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBstore None chunk addr args0 src :: RBexit None (RBcond _ _ _ _) :: nil | RBstore None chunk addr args0 src :: RBexit None (RBjumptable _ _) :: nil | RBstore None chunk addr args0 src :: RBexit None (RBreturn _) :: nil => false | RBstore None chunk addr args0 src :: (RBnop :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBop _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit (Some _) _ :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args0 args'0 && ceq peq src src' && check_code c tc pc''0 b' | RBstore None chunk addr args0 src :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args0 args'0 && ceq peq src src' | RBstore None chunk addr args0 src :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args0 args'0 && ceq peq src src' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Icall sig' (inl r') args'0 dst'0 pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl r) args0 dst0 pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBnop :: RBexit None (RBcall sig (inl r) args0 dst0 pc') :: _ :: _ => false | RBnop :: RBexit None (RBcall sig (inr _) _ _ _) :: _ => false | _ => false end | Some (RTL.Icall sig' (inr i') args'0 dst'0 pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl _) _ _ _) :: _ => false | RBnop :: RBexit None (RBcall sig (inr i) args0 dst0 pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq signature_eq sig sig' && ceq peq i i' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBnop :: RBexit None (RBcall sig (inr i) args0 dst0 pc') :: _ :: _ => false | _ => false end | Some (RTL.Itailcall sig (inl r) args0) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl r') args'0) :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args0 args'0 | RBnop :: RBexit None (RBtailcall sig' (inl r') args'0) :: _ :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr _) _) :: _ => false | _ => false end | Some (RTL.Itailcall sig (inr r) args0) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl _) _) :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr r') args'0) :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args0 args'0 | RBnop :: RBexit None (RBtailcall sig' (inr r') args'0) :: _ :: _ => false | _ => false end | Some (RTL.Icond cond args0 n1 n2) => match b with | RBnop :: RBexit None (RBcond cond' args'0 n1' n2') :: nil => check_valid_node tc n1 && check_valid_node tc n2 && ceq condition_eq cond cond' && ceq list_pos_eq args0 args'0 && ceq peq n1 n1' && ceq peq n2 n2' | RBnop :: RBexit None (RBcond cond' args'0 n1' n2') :: _ :: _ => false | _ => false end | Some (RTL.Ijumptable r ns) => match b with | RBnop :: RBexit None (RBjumptable r' ns') :: nil => ceq peq r r' && ceq list_pos_eq ns ns' && forallb (check_valid_node tc) ns | RBnop :: RBexit None (RBjumptable r' ns') :: _ :: _ => false | _ => false end | Some (RTL.Ireturn (Some r)) => match b with | RBnop :: RBexit None (RBreturn (Some r')) :: nil => ceq peq r r' | RBnop :: RBexit None (RBreturn (Some r')) :: _ :: _ => false | _ => false end | Some (RTL.Ireturn None) => match b with | RBnop :: RBexit None (RBreturn None) :: nil => true | _ => false end | _ => false end) c tc pc'' b | RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc'' && ceq operation_eq op op' && ceq list_pos_eq args args' && ceq peq dst dst' | RBexit None (RBgoto pc') :: _ :: _ => ceq operation_eq op op' && ceq list_pos_eq args args' && ceq peq dst dst' && (fix check_code (c : RTL.code) (tc : code) (pc : node) (b : SeqBB.t) {struct b} : bool := match c ! pc with | Some (RTL.Inop pc''0) => match b with | RBnop :: (RBnop :: _ :: _) as b' | RBnop :: (RBop _ _ _ _ :: _ :: _) as b' | RBnop :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBnop :: (RBexit (Some _) _ :: _ :: _) as b' | RBnop :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBreturn _) :: _ :: _) as b' => check_code c tc pc''0 b' | RBnop :: RBexit None (RBgoto pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 | RBnop :: (RBexit None (RBgoto pc'0) :: _ :: _) as b' => check_code c tc pc''0 b' | _ => false end | Some (RTL.Iop op'0 args'0 dst'0 pc''0) => match b with | RBop None op0 args0 dst0 :: nil | RBop None op0 args0 dst0 :: RBnop :: nil | RBop None op0 args0 dst0 :: RBop _ _ _ _ :: nil | RBop None op0 args0 dst0 :: RBload _ _ _ _ _ :: nil | RBop None op0 args0 dst0 :: RBstore _ _ _ _ _ :: nil | RBop None op0 args0 dst0 :: RBsetpred _ _ _ _ :: nil | RBop None op0 args0 dst0 :: RBexit (Some _) _ :: nil | RBop None op0 args0 dst0 :: RBexit None (RBcall _ _ _ _ _) :: nil | RBop None op0 args0 dst0 :: RBexit None (RBtailcall _ _ _) :: nil | RBop None op0 args0 dst0 :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBop None op0 args0 dst0 :: RBexit None (RBcond _ _ _ _) :: nil | RBop None op0 args0 dst0 :: RBexit None (RBjumptable _ _) :: nil | RBop None op0 args0 dst0 :: RBexit None (RBreturn _) :: nil => false | RBop None op0 args0 dst0 :: (RBnop :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBop _ _ _ _ :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit (Some _) _ :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBop None op0 args0 dst0 :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq operation_eq op0 op'0 && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | RBop None op0 args0 dst0 :: RBexit None (RBgoto pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq operation_eq op0 op'0 && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBop None op0 args0 dst0 :: (RBexit None (RBgoto pc'0) :: _ :: _) as b' => ceq operation_eq op0 op'0 && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | _ => false end | Some (RTL.Iload chunk' addr' args'0 dst'0 pc''0) => match b with | RBload None chunk addr args0 dst0 :: nil | RBload None chunk addr args0 dst0 :: RBnop :: nil | RBload None chunk addr args0 dst0 :: RBop _ _ _ _ :: nil | RBload None chunk addr args0 dst0 :: RBload _ _ _ _ _ :: nil | RBload None chunk addr args0 dst0 :: RBstore _ _ _ _ _ :: nil | RBload None chunk addr args0 dst0 :: RBsetpred _ _ _ _ :: nil | RBload None chunk addr args0 dst0 :: RBexit (Some _) _ :: nil | RBload None chunk addr args0 dst0 :: RBexit None (RBcall _ _ _ _ _) :: nil | RBload None chunk addr args0 dst0 :: RBexit None (RBtailcall _ _ _) :: nil | RBload None chunk addr args0 dst0 :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBload None chunk addr args0 dst0 :: RBexit None (RBcond _ _ _ _) :: nil | RBload None chunk addr args0 dst0 :: RBexit None (RBjumptable _ _) :: nil | RBload None chunk addr args0 dst0 :: RBexit None (RBreturn _) :: nil => false | RBload None chunk addr args0 dst0 :: (RBnop :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBop _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit (Some _) _ :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBload None chunk addr args0 dst0 :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | RBload None chunk addr args0 dst0 :: RBexit None (RBgoto pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBload None chunk addr args0 dst0 :: (RBexit None (RBgoto pc'0) :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | _ => false end | Some (RTL.Istore chunk' addr' args'0 src' pc''0) => match b with | RBstore None chunk addr args0 src :: nil | RBstore None chunk addr args0 src :: RBnop :: nil | RBstore None chunk addr args0 src :: RBop _ _ _ _ :: nil | RBstore None chunk addr args0 src :: RBload _ _ _ _ _ :: nil | RBstore None chunk addr args0 src :: RBstore _ _ _ _ _ :: nil | RBstore None chunk addr args0 src :: RBsetpred _ _ _ _ :: nil | RBstore None chunk addr args0 src :: RBexit (Some _) _ :: nil | RBstore None chunk addr args0 src :: RBexit None (RBcall _ _ _ _ _) :: nil | RBstore None chunk addr args0 src :: RBexit None (RBtailcall _ _ _) :: nil | RBstore None chunk addr args0 src :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBstore None chunk addr args0 src :: RBexit None (RBcond _ _ _ _) :: nil | RBstore None chunk addr args0 src :: RBexit None (RBjumptable _ _) :: nil | RBstore None chunk addr args0 src :: RBexit None (RBreturn _) :: nil => false | RBstore None chunk addr args0 src :: (RBnop :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBop _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit (Some _) _ :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBstore None chunk addr args0 src :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args0 args'0 && ceq peq src src' && check_code c tc pc''0 b' | RBstore None chunk addr args0 src :: RBexit None (RBgoto pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args0 args'0 && ceq peq src src' | RBstore None chunk addr args0 src :: (RBexit None (RBgoto pc'0) :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args0 args'0 && ceq peq src src' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Icall sig' (inl r') args'0 dst'0 pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl r) args0 dst0 pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBnop :: RBexit None (RBcall sig (inl r) args0 dst0 pc'0) :: _ :: _ => false | RBnop :: RBexit None (RBcall sig (inr _) _ _ _) :: _ => false | _ => false end | Some (RTL.Icall sig' (inr i') args'0 dst'0 pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl _) _ _ _) :: _ => false | RBnop :: RBexit None (RBcall sig (inr i) args0 dst0 pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq signature_eq sig sig' && ceq peq i i' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBnop :: RBexit None (RBcall sig (inr i) args0 dst0 pc'0) :: _ :: _ => false | _ => false end | Some (RTL.Itailcall sig (inl r) args0) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl r') args'0) :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args0 args'0 | RBnop :: RBexit None (RBtailcall sig' (inl r') args'0) :: _ :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr _) _) :: _ => false | _ => false end | Some (RTL.Itailcall sig (inr r) args0) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl _) _) :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr r') args'0) :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args0 args'0 | RBnop :: RBexit None (RBtailcall sig' (inr r') args'0) :: _ :: _ => false | _ => false end | Some (RTL.Icond cond args0 n1 n2) => match b with | RBnop :: RBexit None (RBcond cond' args'0 n1' n2') :: nil => check_valid_node tc n1 && check_valid_node tc n2 && ceq condition_eq cond cond' && ceq list_pos_eq args0 args'0 && ceq peq n1 n1' && ceq peq n2 n2' | RBnop :: RBexit None (RBcond cond' args'0 n1' n2') :: _ :: _ => false | _ => false end | Some (RTL.Ijumptable r ns) => match b with | RBnop :: RBexit None (RBjumptable r' ns') :: nil => ceq peq r r' && ceq list_pos_eq ns ns' && forallb (check_valid_node tc) ns | RBnop :: RBexit None (RBjumptable r' ns') :: _ :: _ => false | _ => false end | Some (RTL.Ireturn (Some r)) => match b with | RBnop :: RBexit None (RBreturn (Some r')) :: nil => ceq peq r r' | RBnop :: RBexit None (RBreturn (Some r')) :: _ :: _ => false | _ => false end | Some (RTL.Ireturn None) => match b with | RBnop :: RBexit None (RBreturn None) :: nil => true | _ => false end | _ => false end) c tc pc'' b | _ => false end | _ => false end | Some (RTL.Iload chunk' addr' args' dst' pc'') => match a with | RBload None chunk addr args dst => match b with | RBop _ _ _ _ :: _ :: _ | RBload _ _ _ _ _ :: _ :: _ | RBstore _ _ _ _ _ :: _ :: _ | RBexit None (RBcall _ _ _ _ _) :: _ :: _ | RBexit None (RBjumptable _ _) :: _ :: _ => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq dst dst' && (fix check_code (c : RTL.code) (tc : code) (pc : node) (b : SeqBB.t) {struct b} : bool := match c ! pc with | Some (RTL.Inop pc''0) => match b with | RBnop :: (RBnop :: _ :: _) as b' | RBnop :: (RBop _ _ _ _ :: _ :: _) as b' | RBnop :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBnop :: (RBexit (Some _) _ :: _ :: _) as b' | RBnop :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBreturn _) :: _ :: _) as b' => check_code c tc pc''0 b' | RBnop :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 | RBnop :: (RBexit None (RBgoto pc') :: _ :: _) as b' => check_code c tc pc''0 b' | _ => false end | Some (RTL.Iop op' args'0 dst'0 pc''0) => match b with | RBop None op args0 dst0 :: nil | RBop None op args0 dst0 :: RBnop :: nil | RBop None op args0 dst0 :: RBop _ _ _ _ :: nil | RBop None op args0 dst0 :: RBload _ _ _ _ _ :: nil | RBop None op args0 dst0 :: RBstore _ _ _ _ _ :: nil | RBop None op args0 dst0 :: RBsetpred _ _ _ _ :: nil | RBop None op args0 dst0 :: RBexit (Some _) _ :: nil | RBop None op args0 dst0 :: RBexit None (RBcall _ _ _ _ _) :: nil | RBop None op args0 dst0 :: RBexit None (RBtailcall _ _ _) :: nil | RBop None op args0 dst0 :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBop None op args0 dst0 :: RBexit None (RBcond _ _ _ _) :: nil | RBop None op args0 dst0 :: RBexit None (RBjumptable _ _) :: nil | RBop None op args0 dst0 :: RBexit None (RBreturn _) :: nil => false | RBop None op args0 dst0 :: (RBnop :: _ :: _) as b' | RBop None op args0 dst0 :: (RBop _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst0 :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst0 :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst0 :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit (Some _) _ :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | RBop None op args0 dst0 :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq operation_eq op op' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBop None op args0 dst0 :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | _ => false end | Some (RTL.Iload chunk'0 addr'0 args'0 dst'0 pc''0) => match b with | RBload None chunk0 addr0 args0 dst0 :: nil | RBload None chunk0 addr0 args0 dst0 :: RBnop :: nil | RBload None chunk0 addr0 args0 dst0 :: RBop _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst0 :: RBload _ _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst0 :: RBstore _ _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst0 :: RBsetpred _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit (Some _) _ :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBcall _ _ _ _ _) :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBtailcall _ _ _) :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBcond _ _ _ _) :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBjumptable _ _) :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBreturn _) :: nil => false | RBload None chunk0 addr0 args0 dst0 :: (RBnop :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBop _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit (Some _) _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | _ => false end | Some (RTL.Istore chunk'0 addr'0 args'0 src' pc''0) => match b with | RBstore None chunk0 addr0 args0 src :: nil | RBstore None chunk0 addr0 args0 src :: RBnop :: nil | RBstore None chunk0 addr0 args0 src :: RBop _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src :: RBload _ _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src :: RBstore _ _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src :: RBsetpred _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src :: RBexit (Some _) _ :: nil | RBstore None chunk0 addr0 args0 src :: RBexit None (RBcall _ _ _ _ _) :: nil | RBstore None chunk0 addr0 args0 src :: RBexit None (RBtailcall _ _ _) :: nil | RBstore None chunk0 addr0 args0 src :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBstore None chunk0 addr0 args0 src :: RBexit None (RBcond _ _ _ _) :: nil | RBstore None chunk0 addr0 args0 src :: RBexit None (RBjumptable _ _) :: nil | RBstore None chunk0 addr0 args0 src :: RBexit None (RBreturn _) :: nil => false | RBstore None chunk0 addr0 args0 src :: (RBnop :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBop _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit (Some _) _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq src src' && check_code c tc pc''0 b' | RBstore None chunk0 addr0 args0 src :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq src src' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq src src' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Icall sig' (inl r') args'0 dst'0 pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl r0) args0 dst0 pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq signature_eq sig sig' && ceq peq r0 r' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBnop :: RBexit None (RBcall sig (inl r0) args0 dst0 pc') :: _ :: _ => false | RBnop :: RBexit None (RBcall sig (inr _) _ _ _) :: _ => false | _ => false end | Some (RTL.Icall sig' (inr i') args'0 dst'0 pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl _) _ _ _) :: _ => false | RBnop :: RBexit None (RBcall sig (inr i) args0 dst0 pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq signature_eq sig sig' && ceq peq i i' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBnop :: RBexit None (RBcall sig (inr i) args0 dst0 pc') :: _ :: _ => false | _ => false end | Some (RTL.Itailcall sig (inl r0) args0) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl r') args'0) :: nil => ceq signature_eq sig sig' && ceq peq r0 r' && ceq list_pos_eq args0 args'0 | RBnop :: RBexit None (RBtailcall sig' (inl r') args'0) :: _ :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr _) _) :: _ => false | _ => false end | Some (RTL.Itailcall sig (inr r0) args0) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl _) _) :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr r') args'0) :: nil => ceq signature_eq sig sig' && ceq peq r0 r' && ceq list_pos_eq args0 args'0 | RBnop :: RBexit None (RBtailcall sig' (inr r') args'0) :: _ :: _ => false | _ => false end | Some (RTL.Icond cond args0 n1 n2) => match b with | RBnop :: RBexit None (RBcond cond' args'0 n1' n2') :: nil => check_valid_node tc n1 && check_valid_node tc n2 && ceq condition_eq cond cond' && ceq list_pos_eq args0 args'0 && ceq peq n1 n1' && ceq peq n2 n2' | RBnop :: RBexit None (RBcond cond' args'0 n1' n2') :: _ :: _ => false | _ => false end | Some (RTL.Ijumptable r0 ns) => match b with | RBnop :: RBexit None (RBjumptable r' ns') :: nil => ceq peq r0 r' && ceq list_pos_eq ns ns' && forallb (check_valid_node tc) ns | RBnop :: RBexit None (RBjumptable r' ns') :: _ :: _ => false | _ => false end | Some (RTL.Ireturn (Some r0)) => match b with | RBnop :: RBexit None (RBreturn (Some r')) :: nil => ceq peq r0 r' | RBnop :: RBexit None (RBreturn (Some r')) :: _ :: _ => false | _ => false end | Some (RTL.Ireturn None) => match b with | RBnop :: RBexit None (RBreturn None) :: nil => true | _ => false end | _ => false end) c tc pc'' b | RBnop :: _ :: _ | RBsetpred _ _ _ _ :: _ :: _ | RBexit (Some _) _ :: _ :: _ | RBexit None (RBtailcall _ _ _) :: _ :: _ | RBexit None (RBbuiltin _ _ _ _) :: _ :: _ | RBexit None (RBcond _ _ _ _) :: _ :: _ | RBexit None (RBreturn _) :: _ :: _ => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq dst dst' && (fix check_code (c : RTL.code) (tc : code) (pc : node) (b : SeqBB.t) {struct b} : bool := match c ! pc with | Some (RTL.Inop pc''0) => match b with | RBnop :: (RBnop :: _ :: _) as b' | RBnop :: (RBop _ _ _ _ :: _ :: _) as b' | RBnop :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBnop :: (RBexit (Some _) _ :: _ :: _) as b' | RBnop :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBreturn _) :: _ :: _) as b' => check_code c tc pc''0 b' | RBnop :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 | RBnop :: (RBexit None (RBgoto pc') :: _ :: _) as b' => check_code c tc pc''0 b' | _ => false end | Some (RTL.Iop op' args'0 dst'0 pc''0) => match b with | RBop None op args0 dst0 :: nil | RBop None op args0 dst0 :: RBnop :: nil | RBop None op args0 dst0 :: RBop _ _ _ _ :: nil | RBop None op args0 dst0 :: RBload _ _ _ _ _ :: nil | RBop None op args0 dst0 :: RBstore _ _ _ _ _ :: nil | RBop None op args0 dst0 :: RBsetpred _ _ _ _ :: nil | RBop None op args0 dst0 :: RBexit (Some _) _ :: nil | RBop None op args0 dst0 :: RBexit None (RBcall _ _ _ _ _) :: nil | RBop None op args0 dst0 :: RBexit None (RBtailcall _ _ _) :: nil | RBop None op args0 dst0 :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBop None op args0 dst0 :: RBexit None (RBcond _ _ _ _) :: nil | RBop None op args0 dst0 :: RBexit None (RBjumptable _ _) :: nil | RBop None op args0 dst0 :: RBexit None (RBreturn _) :: nil => false | RBop None op args0 dst0 :: (RBnop :: _ :: _) as b' | RBop None op args0 dst0 :: (RBop _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst0 :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst0 :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst0 :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit (Some _) _ :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | RBop None op args0 dst0 :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq operation_eq op op' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBop None op args0 dst0 :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | _ => false end | Some (RTL.Iload chunk'0 addr'0 args'0 dst'0 pc''0) => match b with | RBload None chunk0 addr0 args0 dst0 :: nil | RBload None chunk0 addr0 args0 dst0 :: RBnop :: nil | RBload None chunk0 addr0 args0 dst0 :: RBop _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst0 :: RBload _ _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst0 :: RBstore _ _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst0 :: RBsetpred _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit (Some _) _ :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBcall _ _ _ _ _) :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBtailcall _ _ _) :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBcond _ _ _ _) :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBjumptable _ _) :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBreturn _) :: nil => false | RBload None chunk0 addr0 args0 dst0 :: (RBnop :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBop _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit (Some _) _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | _ => false end | Some (RTL.Istore chunk'0 addr'0 args'0 src' pc''0) => match b with | RBstore None chunk0 addr0 args0 src :: nil | RBstore None chunk0 addr0 args0 src :: RBnop :: nil | RBstore None chunk0 addr0 args0 src :: RBop _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src :: RBload _ _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src :: RBstore _ _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src :: RBsetpred _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src :: RBexit (Some _) _ :: nil | RBstore None chunk0 addr0 args0 src :: RBexit None (RBcall _ _ _ _ _) :: nil | RBstore None chunk0 addr0 args0 src :: RBexit None (RBtailcall _ _ _) :: nil | RBstore None chunk0 addr0 args0 src :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBstore None chunk0 addr0 args0 src :: RBexit None (RBcond _ _ _ _) :: nil | RBstore None chunk0 addr0 args0 src :: RBexit None (RBjumptable _ _) :: nil | RBstore None chunk0 addr0 args0 src :: RBexit None (RBreturn _) :: nil => false | RBstore None chunk0 addr0 args0 src :: (RBnop :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBop _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit (Some _) _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq src src' && check_code c tc pc''0 b' | RBstore None chunk0 addr0 args0 src :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq src src' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq src src' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Icall sig' (inl r') args'0 dst'0 pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl r) args0 dst0 pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBnop :: RBexit None (RBcall sig (inl r) args0 dst0 pc') :: _ :: _ => false | RBnop :: RBexit None (RBcall sig (inr _) _ _ _) :: _ => false | _ => false end | Some (RTL.Icall sig' (inr i') args'0 dst'0 pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl _) _ _ _) :: _ => false | RBnop :: RBexit None (RBcall sig (inr i) args0 dst0 pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq signature_eq sig sig' && ceq peq i i' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBnop :: RBexit None (RBcall sig (inr i) args0 dst0 pc') :: _ :: _ => false | _ => false end | Some (RTL.Itailcall sig (inl r) args0) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl r') args'0) :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args0 args'0 | RBnop :: RBexit None (RBtailcall sig' (inl r') args'0) :: _ :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr _) _) :: _ => false | _ => false end | Some (RTL.Itailcall sig (inr r) args0) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl _) _) :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr r') args'0) :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args0 args'0 | RBnop :: RBexit None (RBtailcall sig' (inr r') args'0) :: _ :: _ => false | _ => false end | Some (RTL.Icond cond args0 n1 n2) => match b with | RBnop :: RBexit None (RBcond cond' args'0 n1' n2') :: nil => check_valid_node tc n1 && check_valid_node tc n2 && ceq condition_eq cond cond' && ceq list_pos_eq args0 args'0 && ceq peq n1 n1' && ceq peq n2 n2' | RBnop :: RBexit None (RBcond cond' args'0 n1' n2') :: _ :: _ => false | _ => false end | Some (RTL.Ijumptable r ns) => match b with | RBnop :: RBexit None (RBjumptable r' ns') :: nil => ceq peq r r' && ceq list_pos_eq ns ns' && forallb (check_valid_node tc) ns | RBnop :: RBexit None (RBjumptable r' ns') :: _ :: _ => false | _ => false end | Some (RTL.Ireturn (Some r)) => match b with | RBnop :: RBexit None (RBreturn (Some r')) :: nil => ceq peq r r' | RBnop :: RBexit None (RBreturn (Some r')) :: _ :: _ => false | _ => false end | Some (RTL.Ireturn None) => match b with | RBnop :: RBexit None (RBreturn None) :: nil => true | _ => false end | _ => false end) c tc pc'' b | RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc'' && ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq dst dst' | RBexit None (RBgoto pc') :: _ :: _ => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq dst dst' && (fix check_code (c : RTL.code) (tc : code) (pc : node) (b : SeqBB.t) {struct b} : bool := match c ! pc with | Some (RTL.Inop pc''0) => match b with | RBnop :: (RBnop :: _ :: _) as b' | RBnop :: (RBop _ _ _ _ :: _ :: _) as b' | RBnop :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBnop :: (RBexit (Some _) _ :: _ :: _) as b' | RBnop :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBreturn _) :: _ :: _) as b' => check_code c tc pc''0 b' | RBnop :: RBexit None (RBgoto pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 | RBnop :: (RBexit None (RBgoto pc'0) :: _ :: _) as b' => check_code c tc pc''0 b' | _ => false end | Some (RTL.Iop op' args'0 dst'0 pc''0) => match b with | RBop None op args0 dst0 :: nil | RBop None op args0 dst0 :: RBnop :: nil | RBop None op args0 dst0 :: RBop _ _ _ _ :: nil | RBop None op args0 dst0 :: RBload _ _ _ _ _ :: nil | RBop None op args0 dst0 :: RBstore _ _ _ _ _ :: nil | RBop None op args0 dst0 :: RBsetpred _ _ _ _ :: nil | RBop None op args0 dst0 :: RBexit (Some _) _ :: nil | RBop None op args0 dst0 :: RBexit None (RBcall _ _ _ _ _) :: nil | RBop None op args0 dst0 :: RBexit None (RBtailcall _ _ _) :: nil | RBop None op args0 dst0 :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBop None op args0 dst0 :: RBexit None (RBcond _ _ _ _) :: nil | RBop None op args0 dst0 :: RBexit None (RBjumptable _ _) :: nil | RBop None op args0 dst0 :: RBexit None (RBreturn _) :: nil => false | RBop None op args0 dst0 :: (RBnop :: _ :: _) as b' | RBop None op args0 dst0 :: (RBop _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst0 :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst0 :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst0 :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit (Some _) _ :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBop None op args0 dst0 :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | RBop None op args0 dst0 :: RBexit None (RBgoto pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq operation_eq op op' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBop None op args0 dst0 :: (RBexit None (RBgoto pc'0) :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | _ => false end | Some (RTL.Iload chunk'0 addr'0 args'0 dst'0 pc''0) => match b with | RBload None chunk0 addr0 args0 dst0 :: nil | RBload None chunk0 addr0 args0 dst0 :: RBnop :: nil | RBload None chunk0 addr0 args0 dst0 :: RBop _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst0 :: RBload _ _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst0 :: RBstore _ _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst0 :: RBsetpred _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit (Some _) _ :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBcall _ _ _ _ _) :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBtailcall _ _ _) :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBcond _ _ _ _) :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBjumptable _ _) :: nil | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBreturn _) :: nil => false | RBload None chunk0 addr0 args0 dst0 :: (RBnop :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBop _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit (Some _) _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | RBload None chunk0 addr0 args0 dst0 :: RBexit None (RBgoto pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBload None chunk0 addr0 args0 dst0 :: (RBexit None (RBgoto pc'0) :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 && check_code c tc pc''0 b' | _ => false end | Some (RTL.Istore chunk'0 addr'0 args'0 src' pc''0) => match b with | RBstore None chunk0 addr0 args0 src :: nil | RBstore None chunk0 addr0 args0 src :: RBnop :: nil | RBstore None chunk0 addr0 args0 src :: RBop _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src :: RBload _ _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src :: RBstore _ _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src :: RBsetpred _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src :: RBexit (Some _) _ :: nil | RBstore None chunk0 addr0 args0 src :: RBexit None (RBcall _ _ _ _ _) :: nil | RBstore None chunk0 addr0 args0 src :: RBexit None (RBtailcall _ _ _) :: nil | RBstore None chunk0 addr0 args0 src :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBstore None chunk0 addr0 args0 src :: RBexit None (RBcond _ _ _ _) :: nil | RBstore None chunk0 addr0 args0 src :: RBexit None (RBjumptable _ _) :: nil | RBstore None chunk0 addr0 args0 src :: RBexit None (RBreturn _) :: nil => false | RBstore None chunk0 addr0 args0 src :: (RBnop :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBop _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit (Some _) _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq src src' && check_code c tc pc''0 b' | RBstore None chunk0 addr0 args0 src :: RBexit None (RBgoto pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq src src' | RBstore None chunk0 addr0 args0 src :: (RBexit None (RBgoto pc'0) :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq src src' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Icall sig' (inl r') args'0 dst'0 pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl r) args0 dst0 pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBnop :: RBexit None (RBcall sig (inl r) args0 dst0 pc'0) :: _ :: _ => false | RBnop :: RBexit None (RBcall sig (inr _) _ _ _) :: _ => false | _ => false end | Some (RTL.Icall sig' (inr i') args'0 dst'0 pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl _) _ _ _) :: _ => false | RBnop :: RBexit None (RBcall sig (inr i) args0 dst0 pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq signature_eq sig sig' && ceq peq i i' && ceq list_pos_eq args0 args'0 && ceq peq dst0 dst'0 | RBnop :: RBexit None (RBcall sig (inr i) args0 dst0 pc'0) :: _ :: _ => false | _ => false end | Some (RTL.Itailcall sig (inl r) args0) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl r') args'0) :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args0 args'0 | RBnop :: RBexit None (RBtailcall sig' (inl r') args'0) :: _ :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr _) _) :: _ => false | _ => false end | Some (RTL.Itailcall sig (inr r) args0) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl _) _) :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr r') args'0) :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args0 args'0 | RBnop :: RBexit None (RBtailcall sig' (inr r') args'0) :: _ :: _ => false | _ => false end | Some (RTL.Icond cond args0 n1 n2) => match b with | RBnop :: RBexit None (RBcond cond' args'0 n1' n2') :: nil => check_valid_node tc n1 && check_valid_node tc n2 && ceq condition_eq cond cond' && ceq list_pos_eq args0 args'0 && ceq peq n1 n1' && ceq peq n2 n2' | RBnop :: RBexit None (RBcond cond' args'0 n1' n2') :: _ :: _ => false | _ => false end | Some (RTL.Ijumptable r ns) => match b with | RBnop :: RBexit None (RBjumptable r' ns') :: nil => ceq peq r r' && ceq list_pos_eq ns ns' && forallb (check_valid_node tc) ns | RBnop :: RBexit None (RBjumptable r' ns') :: _ :: _ => false | _ => false end | Some (RTL.Ireturn (Some r)) => match b with | RBnop :: RBexit None (RBreturn (Some r')) :: nil => ceq peq r r' | RBnop :: RBexit None (RBreturn (Some r')) :: _ :: _ => false | _ => false end | Some (RTL.Ireturn None) => match b with | RBnop :: RBexit None (RBreturn None) :: nil => true | _ => false end | _ => false end) c tc pc'' b | _ => false end | _ => false end | Some (RTL.Istore chunk' addr' args' src' pc'') => match a with | RBstore None chunk addr args src => match b with | RBop _ _ _ _ :: _ :: _ | RBload _ _ _ _ _ :: _ :: _ | RBstore _ _ _ _ _ :: _ :: _ | RBexit None (RBcall _ _ _ _ _) :: _ :: _ | RBexit None (RBjumptable _ _) :: _ :: _ => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq src src' && (fix check_code (c : RTL.code) (tc : code) (pc : node) (b : SeqBB.t) {struct b} : bool := match c ! pc with | Some (RTL.Inop pc''0) => match b with | RBnop :: (RBnop :: _ :: _) as b' | RBnop :: (RBop _ _ _ _ :: _ :: _) as b' | RBnop :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBnop :: (RBexit (Some _) _ :: _ :: _) as b' | RBnop :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBreturn _) :: _ :: _) as b' => check_code c tc pc''0 b' | RBnop :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 | RBnop :: (RBexit None (RBgoto pc') :: _ :: _) as b' => check_code c tc pc''0 b' | _ => false end | Some (RTL.Iop op' args'0 dst' pc''0) => match b with | RBop None op args0 dst :: nil | RBop None op args0 dst :: RBnop :: nil | RBop None op args0 dst :: RBop _ _ _ _ :: nil | RBop None op args0 dst :: RBload _ _ _ _ _ :: nil | RBop None op args0 dst :: RBstore _ _ _ _ _ :: nil | RBop None op args0 dst :: RBsetpred _ _ _ _ :: nil | RBop None op args0 dst :: RBexit (Some _) _ :: nil | RBop None op args0 dst :: RBexit None (RBcall _ _ _ _ _) :: nil | RBop None op args0 dst :: RBexit None (RBtailcall _ _ _) :: nil | RBop None op args0 dst :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBop None op args0 dst :: RBexit None (RBcond _ _ _ _) :: nil | RBop None op args0 dst :: RBexit None (RBjumptable _ _) :: nil | RBop None op args0 dst :: RBexit None (RBreturn _) :: nil => false | RBop None op args0 dst :: (RBnop :: _ :: _) as b' | RBop None op args0 dst :: (RBop _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst :: (RBexit (Some _) _ :: _ :: _) as b' | RBop None op args0 dst :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBop None op args0 dst :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBop None op args0 dst :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBop None op args0 dst :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBop None op args0 dst :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBop None op args0 dst :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args0 args'0 && ceq peq dst dst' && check_code c tc pc''0 b' | RBop None op args0 dst :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq operation_eq op op' && ceq list_pos_eq args0 args'0 && ceq peq dst dst' | RBop None op args0 dst :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args0 args'0 && ceq peq dst dst' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Iload chunk'0 addr'0 args'0 dst' pc''0) => match b with | RBload None chunk0 addr0 args0 dst :: nil | RBload None chunk0 addr0 args0 dst :: RBnop :: nil | RBload None chunk0 addr0 args0 dst :: RBop _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst :: RBload _ _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst :: RBstore _ _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst :: RBsetpred _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst :: RBexit (Some _) _ :: nil | RBload None chunk0 addr0 args0 dst :: RBexit None (RBcall _ _ _ _ _) :: nil | RBload None chunk0 addr0 args0 dst :: RBexit None (RBtailcall _ _ _) :: nil | RBload None chunk0 addr0 args0 dst :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBload None chunk0 addr0 args0 dst :: RBexit None (RBcond _ _ _ _) :: nil | RBload None chunk0 addr0 args0 dst :: RBexit None (RBjumptable _ _) :: nil | RBload None chunk0 addr0 args0 dst :: RBexit None (RBreturn _) :: nil => false | RBload None chunk0 addr0 args0 dst :: (RBnop :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBop _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit (Some _) _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq dst dst' && check_code c tc pc''0 b' | RBload None chunk0 addr0 args0 dst :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq dst dst' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq dst dst' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Istore chunk'0 addr'0 args'0 src'0 pc''0) => match b with | RBstore None chunk0 addr0 args0 src0 :: nil | RBstore None chunk0 addr0 args0 src0 :: RBnop :: nil | RBstore None chunk0 addr0 args0 src0 :: RBop _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src0 :: RBload _ _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src0 :: RBstore _ _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src0 :: RBsetpred _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit (Some _) _ :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBcall _ _ _ _ _) :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBtailcall _ _ _) :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBcond _ _ _ _) :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBjumptable _ _) :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBreturn _) :: nil => false | RBstore None chunk0 addr0 args0 src0 :: (RBnop :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBop _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit (Some _) _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq src0 src'0 && check_code c tc pc''0 b' | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq src0 src'0 | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq src0 src'0 && check_code c tc pc''0 b' | _ => false end | Some (RTL.Icall sig' (inl r') args'0 dst' pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl r0) args0 dst pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq signature_eq sig sig' && ceq peq r0 r' && ceq list_pos_eq args0 args'0 && ceq peq dst dst' | RBnop :: RBexit None (RBcall sig (inl r0) args0 dst pc') :: _ :: _ => false | RBnop :: RBexit None (RBcall sig (inr _) _ _ _) :: _ => false | _ => false end | Some (RTL.Icall sig' (inr i') args'0 dst' pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl _) _ _ _) :: _ => false | RBnop :: RBexit None (RBcall sig (inr i) args0 dst pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq signature_eq sig sig' && ceq peq i i' && ceq list_pos_eq args0 args'0 && ceq peq dst dst' | RBnop :: RBexit None (RBcall sig (inr i) args0 dst pc') :: _ :: _ => false | _ => false end | Some (RTL.Itailcall sig (inl r0) args0) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl r') args'0) :: nil => ceq signature_eq sig sig' && ceq peq r0 r' && ceq list_pos_eq args0 args'0 | RBnop :: RBexit None (RBtailcall sig' (inl r') args'0) :: _ :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr _) _) :: _ => false | _ => false end | Some (RTL.Itailcall sig (inr r0) args0) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl _) _) :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr r') args'0) :: nil => ceq signature_eq sig sig' && ceq peq r0 r' && ceq list_pos_eq args0 args'0 | RBnop :: RBexit None (RBtailcall sig' (inr r') args'0) :: _ :: _ => false | _ => false end | Some (RTL.Icond cond args0 n1 n2) => match b with | RBnop :: RBexit None (RBcond cond' args'0 n1' n2') :: nil => check_valid_node tc n1 && check_valid_node tc n2 && ceq condition_eq cond cond' && ceq list_pos_eq args0 args'0 && ceq peq n1 n1' && ceq peq n2 n2' | RBnop :: RBexit None (RBcond cond' args'0 n1' n2') :: _ :: _ => false | _ => false end | Some (RTL.Ijumptable r0 ns) => match b with | RBnop :: RBexit None (RBjumptable r' ns') :: nil => ceq peq r0 r' && ceq list_pos_eq ns ns' && forallb (check_valid_node tc) ns | RBnop :: RBexit None (RBjumptable r' ns') :: _ :: _ => false | _ => false end | Some (RTL.Ireturn (Some r0)) => match b with | RBnop :: RBexit None (RBreturn (Some r')) :: nil => ceq peq r0 r' | RBnop :: RBexit None (RBreturn (Some r')) :: _ :: _ => false | _ => false end | Some (RTL.Ireturn None) => match b with | RBnop :: RBexit None (RBreturn None) :: nil => true | _ => false end | _ => false end) c tc pc'' b | RBnop :: _ :: _ | RBsetpred _ _ _ _ :: _ :: _ | RBexit (Some _) _ :: _ :: _ | RBexit None (RBtailcall _ _ _) :: _ :: _ | RBexit None (RBbuiltin _ _ _ _) :: _ :: _ | RBexit None (RBcond _ _ _ _) :: _ :: _ | RBexit None (RBreturn _) :: _ :: _ => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq src src' && (fix check_code (c : RTL.code) (tc : code) (pc : node) (b : SeqBB.t) {struct b} : bool := match c ! pc with | Some (RTL.Inop pc''0) => match b with | RBnop :: (RBnop :: _ :: _) as b' | RBnop :: (RBop _ _ _ _ :: _ :: _) as b' | RBnop :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBnop :: (RBexit (Some _) _ :: _ :: _) as b' | RBnop :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBreturn _) :: _ :: _) as b' => check_code c tc pc''0 b' | RBnop :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 | RBnop :: (RBexit None (RBgoto pc') :: _ :: _) as b' => check_code c tc pc''0 b' | _ => false end | Some (RTL.Iop op' args'0 dst' pc''0) => match b with | RBop None op args0 dst :: nil | RBop None op args0 dst :: RBnop :: nil | RBop None op args0 dst :: RBop _ _ _ _ :: nil | RBop None op args0 dst :: RBload _ _ _ _ _ :: nil | RBop None op args0 dst :: RBstore _ _ _ _ _ :: nil | RBop None op args0 dst :: RBsetpred _ _ _ _ :: nil | RBop None op args0 dst :: RBexit (Some _) _ :: nil | RBop None op args0 dst :: RBexit None (RBcall _ _ _ _ _) :: nil | RBop None op args0 dst :: RBexit None (RBtailcall _ _ _) :: nil | RBop None op args0 dst :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBop None op args0 dst :: RBexit None (RBcond _ _ _ _) :: nil | RBop None op args0 dst :: RBexit None (RBjumptable _ _) :: nil | RBop None op args0 dst :: RBexit None (RBreturn _) :: nil => false | RBop None op args0 dst :: (RBnop :: _ :: _) as b' | RBop None op args0 dst :: (RBop _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst :: (RBexit (Some _) _ :: _ :: _) as b' | RBop None op args0 dst :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBop None op args0 dst :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBop None op args0 dst :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBop None op args0 dst :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBop None op args0 dst :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBop None op args0 dst :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args0 args'0 && ceq peq dst dst' && check_code c tc pc''0 b' | RBop None op args0 dst :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq operation_eq op op' && ceq list_pos_eq args0 args'0 && ceq peq dst dst' | RBop None op args0 dst :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args0 args'0 && ceq peq dst dst' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Iload chunk'0 addr'0 args'0 dst' pc''0) => match b with | RBload None chunk0 addr0 args0 dst :: nil | RBload None chunk0 addr0 args0 dst :: RBnop :: nil | RBload None chunk0 addr0 args0 dst :: RBop _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst :: RBload _ _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst :: RBstore _ _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst :: RBsetpred _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst :: RBexit (Some _) _ :: nil | RBload None chunk0 addr0 args0 dst :: RBexit None (RBcall _ _ _ _ _) :: nil | RBload None chunk0 addr0 args0 dst :: RBexit None (RBtailcall _ _ _) :: nil | RBload None chunk0 addr0 args0 dst :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBload None chunk0 addr0 args0 dst :: RBexit None (RBcond _ _ _ _) :: nil | RBload None chunk0 addr0 args0 dst :: RBexit None (RBjumptable _ _) :: nil | RBload None chunk0 addr0 args0 dst :: RBexit None (RBreturn _) :: nil => false | RBload None chunk0 addr0 args0 dst :: (RBnop :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBop _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit (Some _) _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq dst dst' && check_code c tc pc''0 b' | RBload None chunk0 addr0 args0 dst :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq dst dst' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq dst dst' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Istore chunk'0 addr'0 args'0 src'0 pc''0) => match b with | RBstore None chunk0 addr0 args0 src0 :: nil | RBstore None chunk0 addr0 args0 src0 :: RBnop :: nil | RBstore None chunk0 addr0 args0 src0 :: RBop _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src0 :: RBload _ _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src0 :: RBstore _ _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src0 :: RBsetpred _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit (Some _) _ :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBcall _ _ _ _ _) :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBtailcall _ _ _) :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBcond _ _ _ _) :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBjumptable _ _) :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBreturn _) :: nil => false | RBstore None chunk0 addr0 args0 src0 :: (RBnop :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBop _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit (Some _) _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq src0 src'0 && check_code c tc pc''0 b' | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq src0 src'0 | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq src0 src'0 && check_code c tc pc''0 b' | _ => false end | Some (RTL.Icall sig' (inl r') args'0 dst' pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl r) args0 dst pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args0 args'0 && ceq peq dst dst' | RBnop :: RBexit None (RBcall sig (inl r) args0 dst pc') :: _ :: _ => false | RBnop :: RBexit None (RBcall sig (inr _) _ _ _) :: _ => false | _ => false end | Some (RTL.Icall sig' (inr i') args'0 dst' pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl _) _ _ _) :: _ => false | RBnop :: RBexit None (RBcall sig (inr i) args0 dst pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc''0 && ceq signature_eq sig sig' && ceq peq i i' && ceq list_pos_eq args0 args'0 && ceq peq dst dst' | RBnop :: RBexit None (RBcall sig (inr i) args0 dst pc') :: _ :: _ => false | _ => false end | Some (RTL.Itailcall sig (inl r) args0) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl r') args'0) :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args0 args'0 | RBnop :: RBexit None (RBtailcall sig' (inl r') args'0) :: _ :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr _) _) :: _ => false | _ => false end | Some (RTL.Itailcall sig (inr r) args0) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl _) _) :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr r') args'0) :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args0 args'0 | RBnop :: RBexit None (RBtailcall sig' (inr r') args'0) :: _ :: _ => false | _ => false end | Some (RTL.Icond cond args0 n1 n2) => match b with | RBnop :: RBexit None (RBcond cond' args'0 n1' n2') :: nil => check_valid_node tc n1 && check_valid_node tc n2 && ceq condition_eq cond cond' && ceq list_pos_eq args0 args'0 && ceq peq n1 n1' && ceq peq n2 n2' | RBnop :: RBexit None (RBcond cond' args'0 n1' n2') :: _ :: _ => false | _ => false end | Some (RTL.Ijumptable r ns) => match b with | RBnop :: RBexit None (RBjumptable r' ns') :: nil => ceq peq r r' && ceq list_pos_eq ns ns' && forallb (check_valid_node tc) ns | RBnop :: RBexit None (RBjumptable r' ns') :: _ :: _ => false | _ => false end | Some (RTL.Ireturn (Some r)) => match b with | RBnop :: RBexit None (RBreturn (Some r')) :: nil => ceq peq r r' | RBnop :: RBexit None (RBreturn (Some r')) :: _ :: _ => false | _ => false end | Some (RTL.Ireturn None) => match b with | RBnop :: RBexit None (RBreturn None) :: nil => true | _ => false end | _ => false end) c tc pc'' b | RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc'' && ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq src src' | RBexit None (RBgoto pc') :: _ :: _ => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq src src' && (fix check_code (c : RTL.code) (tc : code) (pc : node) (b : SeqBB.t) {struct b} : bool := match c ! pc with | Some (RTL.Inop pc''0) => match b with | RBnop :: (RBnop :: _ :: _) as b' | RBnop :: (RBop _ _ _ _ :: _ :: _) as b' | RBnop :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBnop :: (RBexit (Some _) _ :: _ :: _) as b' | RBnop :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBreturn _) :: _ :: _) as b' => check_code c tc pc''0 b' | RBnop :: RBexit None (RBgoto pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 | RBnop :: (RBexit None (RBgoto pc'0) :: _ :: _) as b' => check_code c tc pc''0 b' | _ => false end | Some (RTL.Iop op' args'0 dst' pc''0) => match b with | RBop None op args0 dst :: nil | RBop None op args0 dst :: RBnop :: nil | RBop None op args0 dst :: RBop _ _ _ _ :: nil | RBop None op args0 dst :: RBload _ _ _ _ _ :: nil | RBop None op args0 dst :: RBstore _ _ _ _ _ :: nil | RBop None op args0 dst :: RBsetpred _ _ _ _ :: nil | RBop None op args0 dst :: RBexit (Some _) _ :: nil | RBop None op args0 dst :: RBexit None (RBcall _ _ _ _ _) :: nil | RBop None op args0 dst :: RBexit None (RBtailcall _ _ _) :: nil | RBop None op args0 dst :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBop None op args0 dst :: RBexit None (RBcond _ _ _ _) :: nil | RBop None op args0 dst :: RBexit None (RBjumptable _ _) :: nil | RBop None op args0 dst :: RBexit None (RBreturn _) :: nil => false | RBop None op args0 dst :: (RBnop :: _ :: _) as b' | RBop None op args0 dst :: (RBop _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBop None op args0 dst :: (RBexit (Some _) _ :: _ :: _) as b' | RBop None op args0 dst :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBop None op args0 dst :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBop None op args0 dst :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBop None op args0 dst :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBop None op args0 dst :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBop None op args0 dst :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args0 args'0 && ceq peq dst dst' && check_code c tc pc''0 b' | RBop None op args0 dst :: RBexit None (RBgoto pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq operation_eq op op' && ceq list_pos_eq args0 args'0 && ceq peq dst dst' | RBop None op args0 dst :: (RBexit None (RBgoto pc'0) :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args0 args'0 && ceq peq dst dst' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Iload chunk'0 addr'0 args'0 dst' pc''0) => match b with | RBload None chunk0 addr0 args0 dst :: nil | RBload None chunk0 addr0 args0 dst :: RBnop :: nil | RBload None chunk0 addr0 args0 dst :: RBop _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst :: RBload _ _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst :: RBstore _ _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst :: RBsetpred _ _ _ _ :: nil | RBload None chunk0 addr0 args0 dst :: RBexit (Some _) _ :: nil | RBload None chunk0 addr0 args0 dst :: RBexit None (RBcall _ _ _ _ _) :: nil | RBload None chunk0 addr0 args0 dst :: RBexit None (RBtailcall _ _ _) :: nil | RBload None chunk0 addr0 args0 dst :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBload None chunk0 addr0 args0 dst :: RBexit None (RBcond _ _ _ _) :: nil | RBload None chunk0 addr0 args0 dst :: RBexit None (RBjumptable _ _) :: nil | RBload None chunk0 addr0 args0 dst :: RBexit None (RBreturn _) :: nil => false | RBload None chunk0 addr0 args0 dst :: (RBnop :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBop _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit (Some _) _ :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq dst dst' && check_code c tc pc''0 b' | RBload None chunk0 addr0 args0 dst :: RBexit None (RBgoto pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq dst dst' | RBload None chunk0 addr0 args0 dst :: (RBexit None (RBgoto pc'0) :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq dst dst' && check_code c tc pc''0 b' | _ => false end | Some (RTL.Istore chunk'0 addr'0 args'0 src'0 pc''0) => match b with | RBstore None chunk0 addr0 args0 src0 :: nil | RBstore None chunk0 addr0 args0 src0 :: RBnop :: nil | RBstore None chunk0 addr0 args0 src0 :: RBop _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src0 :: RBload _ _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src0 :: RBstore _ _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src0 :: RBsetpred _ _ _ _ :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit (Some _) _ :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBcall _ _ _ _ _) :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBtailcall _ _ _) :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBcond _ _ _ _) :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBjumptable _ _) :: nil | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBreturn _) :: nil => false | RBstore None chunk0 addr0 args0 src0 :: (RBnop :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBop _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit (Some _) _ :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq src0 src'0 && check_code c tc pc''0 b' | RBstore None chunk0 addr0 args0 src0 :: RBexit None (RBgoto pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq src0 src'0 | RBstore None chunk0 addr0 args0 src0 :: (RBexit None (RBgoto pc'0) :: _ :: _) as b' => ceq memory_chunk_eq chunk0 chunk'0 && ceq addressing_eq addr0 addr'0 && ceq list_pos_eq args0 args'0 && ceq peq src0 src'0 && check_code c tc pc''0 b' | _ => false end | Some (RTL.Icall sig' (inl r') args'0 dst' pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl r) args0 dst pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args0 args'0 && ceq peq dst dst' | RBnop :: RBexit None (RBcall sig (inl r) args0 dst pc'0) :: _ :: _ => false | RBnop :: RBexit None (RBcall sig (inr _) _ _ _) :: _ => false | _ => false end | Some (RTL.Icall sig' (inr i') args'0 dst' pc''0) => match b with | RBnop :: RBexit None (RBcall sig (inl _) _ _ _) :: _ => false | RBnop :: RBexit None (RBcall sig (inr i) args0 dst pc'0) :: nil => check_valid_node tc pc'0 && ceq peq pc'0 pc''0 && ceq signature_eq sig sig' && ceq peq i i' && ceq list_pos_eq args0 args'0 && ceq peq dst dst' | RBnop :: RBexit None (RBcall sig (inr i) args0 dst pc'0) :: _ :: _ => false | _ => false end | Some (RTL.Itailcall sig (inl r) args0) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl r') args'0) :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args0 args'0 | RBnop :: RBexit None (RBtailcall sig' (inl r') args'0) :: _ :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr _) _) :: _ => false | _ => false end | Some (RTL.Itailcall sig (inr r) args0) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl _) _) :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr r') args'0) :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args0 args'0 | RBnop :: RBexit None (RBtailcall sig' (inr r') args'0) :: _ :: _ => false | _ => false end | Some (RTL.Icond cond args0 n1 n2) => match b with | RBnop :: RBexit None (RBcond cond' args'0 n1' n2') :: nil => check_valid_node tc n1 && check_valid_node tc n2 && ceq condition_eq cond cond' && ceq list_pos_eq args0 args'0 && ceq peq n1 n1' && ceq peq n2 n2' | RBnop :: RBexit None (RBcond cond' args'0 n1' n2') :: _ :: _ => false | _ => false end | Some (RTL.Ijumptable r ns) => match b with | RBnop :: RBexit None (RBjumptable r' ns') :: nil => ceq peq r r' && ceq list_pos_eq ns ns' && forallb (check_valid_node tc) ns | RBnop :: RBexit None (RBjumptable r' ns') :: _ :: _ => false | _ => false end | Some (RTL.Ireturn (Some r)) => match b with | RBnop :: RBexit None (RBreturn (Some r')) :: nil => ceq peq r r' | RBnop :: RBexit None (RBreturn (Some r')) :: _ :: _ => false | _ => false end | Some (RTL.Ireturn None) => match b with | RBnop :: RBexit None (RBreturn None) :: nil => true | _ => false end | _ => false end) c tc pc'' b | _ => false end | _ => false end | Some (RTL.Icall sig' (inl r') args' dst' pc'') => match a with | RBnop => match b with | RBexit None (RBcall sig (inl r) args dst pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc'' && ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args args' && ceq peq dst dst' | RBexit None (RBcall sig (inl r) args dst pc') :: _ :: _ => false | RBexit None (RBcall sig (inr _) _ _ _) :: _ => false | _ => false end | _ => false end | Some (RTL.Icall sig' (inr i') args' dst' pc'') => match a with | RBnop => match b with | RBexit None (RBcall sig (inl _) _ _ _) :: _ => false | RBexit None (RBcall sig (inr i) args dst pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc'' && ceq signature_eq sig sig' && ceq peq i i' && ceq list_pos_eq args args' && ceq peq dst dst' | RBexit None (RBcall sig (inr i) args dst pc') :: _ :: _ => false | _ => false end | _ => false end | Some (RTL.Itailcall sig (inl r) args) => match a with | RBnop => match b with | RBexit None (RBtailcall sig' (inl r') args') :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args args' | RBexit None (RBtailcall sig' (inl r') args') :: _ :: _ => false | RBexit None (RBtailcall sig' (inr _) _) :: _ => false | _ => false end | _ => false end | Some (RTL.Itailcall sig (inr r) args) => match a with | RBnop => match b with | RBexit None (RBtailcall sig' (inl _) _) :: _ => false | RBexit None (RBtailcall sig' (inr r') args') :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args args' | RBexit None (RBtailcall sig' (inr r') args') :: _ :: _ => false | _ => false end | _ => false end | Some (RTL.Icond cond args n1 n2) => match a with | RBnop => match b with | RBexit None (RBcond cond' args' n1' n2') :: nil => check_valid_node tc n1 && check_valid_node tc n2 && ceq condition_eq cond cond' && ceq list_pos_eq args args' && ceq peq n1 n1' && ceq peq n2 n2' | RBexit None (RBcond cond' args' n1' n2') :: _ :: _ => false | _ => false end | _ => false end | Some (RTL.Ijumptable r ns) => match a with | RBnop => match b with | RBexit None (RBjumptable r' ns') :: nil => ceq peq r r' && ceq list_pos_eq ns ns' && forallb (check_valid_node tc) ns | RBexit None (RBjumptable r' ns') :: _ :: _ => false | _ => false end | _ => false end | Some (RTL.Ireturn (Some r)) => match a with | RBnop => match b with | RBexit None (RBreturn (Some r')) :: nil => ceq peq r r' | RBexit None (RBreturn (Some r')) :: _ :: _ => false | _ => false end | _ => false end | Some (RTL.Ireturn None) => match a with | RBnop => match b with | RBexit None (RBreturn None) :: nil => true | _ => false end | _ => false end | _ => false end = true

match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
n: RTL.node
Heqo: c ! pc = Some (RTL.Inop n)
H: match a with | RBnop => match b with | RBnop :: _ :: _ | RBop _ _ _ _ :: _ :: _ | RBload _ _ _ _ _ :: _ :: _ | RBstore _ _ _ _ _ :: _ :: _ | RBsetpred _ _ _ _ :: _ :: _ | RBexit (Some _) _ :: _ :: _ | RBexit None (RBcall _ _ _ _ _) :: _ :: _ | RBexit None (RBtailcall _ _ _) :: _ :: _ | RBexit None (RBbuiltin _ _ _ _) :: _ :: _ | RBexit None (RBcond _ _ _ _) :: _ :: _ | RBexit None (RBjumptable _ _) :: _ :: _ | RBexit None (RBreturn _) :: _ :: _ => check_code c tc n b | RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' n | RBexit None (RBgoto pc') :: _ :: _ => check_code c tc n b | _ => false end | _ => false end = true

match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
o: Op.operation
l: list reg
r: reg
n: RTL.node
Heqo: c ! pc = Some (RTL.Iop o l r n)
H: match a with | RBop None op args dst => match b with | RBnop :: _ :: _ | RBop _ _ _ _ :: _ :: _ | RBload _ _ _ _ _ :: _ :: _ | RBstore _ _ _ _ _ :: _ :: _ | RBsetpred _ _ _ _ :: _ :: _ | RBexit (Some _) _ :: _ :: _ | RBexit None (RBcall _ _ _ _ _) :: _ :: _ | RBexit None (RBtailcall _ _ _) :: _ :: _ | RBexit None (RBbuiltin _ _ _ _) :: _ :: _ | RBexit None (RBcond _ _ _ _) :: _ :: _ | RBexit None (RBjumptable _ _) :: _ :: _ | RBexit None (RBreturn _) :: _ :: _ => ceq operation_eq op o && ceq list_pos_eq args l && ceq peq dst r && check_code c tc n b | RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' n && ceq operation_eq op o && ceq list_pos_eq args l && ceq peq dst r | RBexit None (RBgoto pc') :: _ :: _ => ceq operation_eq op o && ceq list_pos_eq args l && ceq peq dst r && check_code c tc n b | _ => false end | _ => false end = true
match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
m: memory_chunk
a0: Op.addressing
l: list reg
r: reg
n: RTL.node
Heqo: c ! pc = Some (RTL.Iload m a0 l r n)
H: match a with | RBload None chunk addr args dst => match b with | RBnop :: _ :: _ | RBop _ _ _ _ :: _ :: _ | RBload _ _ _ _ _ :: _ :: _ | RBstore _ _ _ _ _ :: _ :: _ | RBsetpred _ _ _ _ :: _ :: _ | RBexit (Some _) _ :: _ :: _ | RBexit None (RBcall _ _ _ _ _) :: _ :: _ | RBexit None (RBtailcall _ _ _) :: _ :: _ | RBexit None (RBbuiltin _ _ _ _) :: _ :: _ | RBexit None (RBcond _ _ _ _) :: _ :: _ | RBexit None (RBjumptable _ _) :: _ :: _ | RBexit None (RBreturn _) :: _ :: _ => ceq memory_chunk_eq chunk m && ceq addressing_eq addr a0 && ceq list_pos_eq args l && ceq peq dst r && check_code c tc n b | RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' n && ceq memory_chunk_eq chunk m && ceq addressing_eq addr a0 && ceq list_pos_eq args l && ceq peq dst r | RBexit None (RBgoto pc') :: _ :: _ => ceq memory_chunk_eq chunk m && ceq addressing_eq addr a0 && ceq list_pos_eq args l && ceq peq dst r && check_code c tc n b | _ => false end | _ => false end = true
match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
m: memory_chunk
a0: Op.addressing
l: list reg
r: reg
n: RTL.node
Heqo: c ! pc = Some (RTL.Istore m a0 l r n)
H: match a with | RBstore None chunk addr args src => match b with | RBnop :: _ :: _ | RBop _ _ _ _ :: _ :: _ | RBload _ _ _ _ _ :: _ :: _ | RBstore _ _ _ _ _ :: _ :: _ | RBsetpred _ _ _ _ :: _ :: _ | RBexit (Some _) _ :: _ :: _ | RBexit None (RBcall _ _ _ _ _) :: _ :: _ | RBexit None (RBtailcall _ _ _) :: _ :: _ | RBexit None (RBbuiltin _ _ _ _) :: _ :: _ | RBexit None (RBcond _ _ _ _) :: _ :: _ | RBexit None (RBjumptable _ _) :: _ :: _ | RBexit None (RBreturn _) :: _ :: _ => ceq memory_chunk_eq chunk m && ceq addressing_eq addr a0 && ceq list_pos_eq args l && ceq peq src r && check_code c tc n b | RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' n && ceq memory_chunk_eq chunk m && ceq addressing_eq addr a0 && ceq list_pos_eq args l && ceq peq src r | RBexit None (RBgoto pc') :: _ :: _ => ceq memory_chunk_eq chunk m && ceq addressing_eq addr a0 && ceq list_pos_eq args l && ceq peq src r && check_code c tc n b | _ => false end | _ => false end = true
match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
s: signature
s0: (reg + ident)%type
l: list reg
r: reg
n: RTL.node
Heqo: c ! pc = Some (RTL.Icall s s0 l r n)
H: match s0 with | inl r' => match a with | RBnop => match b with | RBexit None (RBcall sig (inl r0) args dst pc') :: nil => check_valid_node tc pc' && ceq peq pc' n && ceq signature_eq sig s && ceq peq r0 r' && ceq list_pos_eq args l && ceq peq dst r | RBexit None (RBcall sig (inl r0) args dst pc') :: _ :: _ => false | RBexit None (RBcall sig (inr _) _ _ _) :: _ => false | _ => false end | _ => false end | inr i' => match a with | RBnop => match b with | RBexit None (RBcall sig (inl _) _ _ _) :: _ => false | RBexit None (RBcall sig (inr i) args dst pc') :: nil => check_valid_node tc pc' && ceq peq pc' n && ceq signature_eq sig s && ceq peq i i' && ceq list_pos_eq args l && ceq peq dst r | RBexit None (RBcall sig (inr i) args dst pc') :: _ :: _ => false | _ => false end | _ => false end end = true
match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
s: signature
s0: (reg + ident)%type
l: list reg
Heqo: c ! pc = Some (RTL.Itailcall s s0 l)
H: match s0 with | inl r => match a with | RBnop => match b with | RBexit None (RBtailcall sig' (inl r') args') :: nil => ceq signature_eq s sig' && ceq peq r r' && ceq list_pos_eq l args' | RBexit None (RBtailcall sig' (inl r') args') :: _ :: _ => false | RBexit None (RBtailcall sig' (inr _) _) :: _ => false | _ => false end | _ => false end | inr r => match a with | RBnop => match b with | RBexit None (RBtailcall sig' (inl _) _) :: _ => false | RBexit None (RBtailcall sig' (inr r') args') :: nil => ceq signature_eq s sig' && ceq peq r r' && ceq list_pos_eq l args' | RBexit None (RBtailcall sig' (inr r') args') :: _ :: _ => false | _ => false end | _ => false end end = true
match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
c0: Op.condition
l: list reg
n, n0: RTL.node
Heqo: c ! pc = Some (RTL.Icond c0 l n n0)
H: match a with | RBnop => match b with | RBexit None (RBcond cond' args' n1' n2') :: nil => check_valid_node tc n && check_valid_node tc n0 && ceq condition_eq c0 cond' && ceq list_pos_eq l args' && ceq peq n n1' && ceq peq n0 n2' | RBexit None (RBcond cond' args' n1' n2') :: _ :: _ => false | _ => false end | _ => false end = true
match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
r: reg
l: list RTL.node
Heqo: c ! pc = Some (RTL.Ijumptable r l)
H: match a with | RBnop => match b with | RBexit None (RBjumptable r' ns') :: nil => ceq peq r r' && ceq list_pos_eq l ns' && forallb (check_valid_node tc) l | RBexit None (RBjumptable r' ns') :: _ :: _ => false | _ => false end | _ => false end = true
match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
o: option reg
Heqo: c ! pc = Some (RTL.Ireturn o)
H: match o with | Some r => match a with | RBnop => match b with | RBexit None (RBreturn (Some r')) :: nil => ceq peq r r' | RBexit None (RBreturn (Some r')) :: _ :: _ => false | _ => false end | _ => false end | None => match a with | RBnop => match b with | RBexit None (RBreturn None) :: nil => true | _ => false end | _ => false end end = true
match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
n: RTL.node
Heqo: c ! pc = Some (RTL.Inop n)
H: match a with | RBnop => match b with | RBnop :: _ :: _ | RBop _ _ _ _ :: _ :: _ | RBload _ _ _ _ _ :: _ :: _ | RBstore _ _ _ _ _ :: _ :: _ | RBsetpred _ _ _ _ :: _ :: _ | RBexit (Some _) _ :: _ :: _ | RBexit None (RBcall _ _ _ _ _) :: _ :: _ | RBexit None (RBtailcall _ _ _) :: _ :: _ | RBexit None (RBbuiltin _ _ _ _) :: _ :: _ | RBexit None (RBcond _ _ _ _) :: _ :: _ | RBexit None (RBjumptable _ _) :: _ :: _ | RBexit None (RBreturn _) :: _ :: _ => check_code c tc n b | RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' n | RBexit None (RBgoto pc') :: _ :: _ => check_code c tc n b | _ => false end | _ => false end = true

match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
i: instr
l: list instr
o: option pred_op
c0: cf_instr
Heqo0: o = None
n0: node
Heqc1: c0 = RBgoto n0
Heqi0: i = RBexit None (RBgoto n0)
Heql0: l = nil
Heql: b = RBexit None (RBgoto n0) :: nil
IHb: forall pc : node, check_code c tc pc (RBexit None (RBgoto n0) :: nil) = true -> match_block c tc pc (RBexit None (RBgoto n0) :: nil)
pc: node
n: RTL.node
Heqo: c ! pc = Some (RTL.Inop n)
Heqi: a = RBnop
H: check_valid_node tc n0 && ceq peq n0 n = true

match_block c tc pc (RBnop :: RBexit None (RBgoto n0) :: nil)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
n: RTL.node
IHb: forall pc : node, check_code c tc pc (RBexit None (RBgoto n) :: nil) = true -> match_block c tc pc (RBexit None (RBgoto n) :: nil)
pc: node
Heqo: c ! pc = Some (RTL.Inop n)
H0: check_valid_node tc n = true

match_block c tc pc (RBnop :: RBexit None (RBgoto n) :: nil)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
n: RTL.node
IHb: forall pc : node, check_code c tc pc (RBexit None (RBgoto n) :: nil) = true -> match_block c tc pc (RBexit None (RBgoto n) :: nil)
pc: node
Heqo: c ! pc = Some (RTL.Inop n)
H0: check_valid_node tc n = true

valid_succ tc n
eauto using check_valid_node_correct.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
o: Op.operation
l: list reg
r: reg
n: RTL.node
Heqo: c ! pc = Some (RTL.Iop o l r n)
H: match a with | RBop None op args dst => match b with | RBnop :: _ :: _ | RBop _ _ _ _ :: _ :: _ | RBload _ _ _ _ _ :: _ :: _ | RBstore _ _ _ _ _ :: _ :: _ | RBsetpred _ _ _ _ :: _ :: _ | RBexit (Some _) _ :: _ :: _ | RBexit None (RBcall _ _ _ _ _) :: _ :: _ | RBexit None (RBtailcall _ _ _) :: _ :: _ | RBexit None (RBbuiltin _ _ _ _) :: _ :: _ | RBexit None (RBcond _ _ _ _) :: _ :: _ | RBexit None (RBjumptable _ _) :: _ :: _ | RBexit None (RBreturn _) :: _ :: _ => ceq operation_eq op o && ceq list_pos_eq args l && ceq peq dst r && check_code c tc n b | RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' n && ceq operation_eq op o && ceq list_pos_eq args l && ceq peq dst r | RBexit None (RBgoto pc') :: _ :: _ => ceq operation_eq op o && ceq list_pos_eq args l && ceq peq dst r && check_code c tc n b | _ => false end | _ => false end = true

match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
n0: node
IHb: forall pc : node, check_code c tc pc (RBexit None (RBgoto n0) :: nil) = true -> match_block c tc pc (RBexit None (RBgoto n0) :: nil)
pc: node
o: Op.operation
l: list reg
r: reg
n: RTL.node
Heqo: c ! pc = Some (RTL.Iop o l r n)
o1: Op.operation
l0: list reg
r0: reg
H: check_valid_node tc n0 && ceq peq n0 n && ceq operation_eq o1 o && ceq list_pos_eq l0 l && ceq peq r0 r = true

match_block c tc pc (RBop None o1 l0 r0 :: RBexit None (RBgoto n0) :: nil)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
n: RTL.node
IHb: forall pc : node, check_code c tc pc (RBexit None (RBgoto n) :: nil) = true -> match_block c tc pc (RBexit None (RBgoto n) :: nil)
pc: node
o: Op.operation
l: list reg
r: reg
Heqo: c ! pc = Some (RTL.Iop o l r n)
H: check_valid_node tc n = true

match_block c tc pc (RBop None o l r :: RBexit None (RBgoto n) :: nil)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
n: RTL.node
IHb: forall pc : node, check_code c tc pc (RBexit None (RBgoto n) :: nil) = true -> match_block c tc pc (RBexit None (RBgoto n) :: nil)
pc: node
o: Op.operation
l: list reg
r: reg
Heqo: c ! pc = Some (RTL.Iop o l r n)
H: check_valid_node tc n = true

valid_succ tc n
eauto using check_valid_node_correct.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
m: memory_chunk
a0: Op.addressing
l: list reg
r: reg
n: RTL.node
Heqo: c ! pc = Some (RTL.Iload m a0 l r n)
H: match a with | RBload None chunk addr args dst => match b with | RBnop :: _ :: _ | RBop _ _ _ _ :: _ :: _ | RBload _ _ _ _ _ :: _ :: _ | RBstore _ _ _ _ _ :: _ :: _ | RBsetpred _ _ _ _ :: _ :: _ | RBexit (Some _) _ :: _ :: _ | RBexit None (RBcall _ _ _ _ _) :: _ :: _ | RBexit None (RBtailcall _ _ _) :: _ :: _ | RBexit None (RBbuiltin _ _ _ _) :: _ :: _ | RBexit None (RBcond _ _ _ _) :: _ :: _ | RBexit None (RBjumptable _ _) :: _ :: _ | RBexit None (RBreturn _) :: _ :: _ => ceq memory_chunk_eq chunk m && ceq addressing_eq addr a0 && ceq list_pos_eq args l && ceq peq dst r && check_code c tc n b | RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' n && ceq memory_chunk_eq chunk m && ceq addressing_eq addr a0 && ceq list_pos_eq args l && ceq peq dst r | RBexit None (RBgoto pc') :: _ :: _ => ceq memory_chunk_eq chunk m && ceq addressing_eq addr a0 && ceq list_pos_eq args l && ceq peq dst r && check_code c tc n b | _ => false end | _ => false end = true

match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
n0: node
IHb: forall pc : node, check_code c tc pc (RBexit None (RBgoto n0) :: nil) = true -> match_block c tc pc (RBexit None (RBgoto n0) :: nil)
pc: node
m: memory_chunk
a0: Op.addressing
l: list reg
r: reg
n: RTL.node
Heqo: c ! pc = Some (RTL.Iload m a0 l r n)
m0: memory_chunk
a1: Op.addressing
l0: list reg
r0: reg
H: check_valid_node tc n0 && ceq peq n0 n && ceq memory_chunk_eq m0 m && ceq addressing_eq a1 a0 && ceq list_pos_eq l0 l && ceq peq r0 r = true

match_block c tc pc (RBload None m0 a1 l0 r0 :: RBexit None (RBgoto n0) :: nil)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
n: RTL.node
IHb: forall pc : node, check_code c tc pc (RBexit None (RBgoto n) :: nil) = true -> match_block c tc pc (RBexit None (RBgoto n) :: nil)
pc: node
m: memory_chunk
a0: Op.addressing
l: list reg
r: reg
Heqo: c ! pc = Some (RTL.Iload m a0 l r n)
H0: check_valid_node tc n = true

match_block c tc pc (RBload None m a0 l r :: RBexit None (RBgoto n) :: nil)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
n: RTL.node
IHb: forall pc : node, check_code c tc pc (RBexit None (RBgoto n) :: nil) = true -> match_block c tc pc (RBexit None (RBgoto n) :: nil)
pc: node
m: memory_chunk
a0: Op.addressing
l: list reg
r: reg
Heqo: c ! pc = Some (RTL.Iload m a0 l r n)
H0: check_valid_node tc n = true

valid_succ tc n
eauto using check_valid_node_correct.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
m: memory_chunk
a0: Op.addressing
l: list reg
r: reg
n: RTL.node
Heqo: c ! pc = Some (RTL.Istore m a0 l r n)
H: match a with | RBstore None chunk addr args src => match b with | RBnop :: _ :: _ | RBop _ _ _ _ :: _ :: _ | RBload _ _ _ _ _ :: _ :: _ | RBstore _ _ _ _ _ :: _ :: _ | RBsetpred _ _ _ _ :: _ :: _ | RBexit (Some _) _ :: _ :: _ | RBexit None (RBcall _ _ _ _ _) :: _ :: _ | RBexit None (RBtailcall _ _ _) :: _ :: _ | RBexit None (RBbuiltin _ _ _ _) :: _ :: _ | RBexit None (RBcond _ _ _ _) :: _ :: _ | RBexit None (RBjumptable _ _) :: _ :: _ | RBexit None (RBreturn _) :: _ :: _ => ceq memory_chunk_eq chunk m && ceq addressing_eq addr a0 && ceq list_pos_eq args l && ceq peq src r && check_code c tc n b | RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' n && ceq memory_chunk_eq chunk m && ceq addressing_eq addr a0 && ceq list_pos_eq args l && ceq peq src r | RBexit None (RBgoto pc') :: _ :: _ => ceq memory_chunk_eq chunk m && ceq addressing_eq addr a0 && ceq list_pos_eq args l && ceq peq src r && check_code c tc n b | _ => false end | _ => false end = true

match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
n0: node
IHb: forall pc : node, check_code c tc pc (RBexit None (RBgoto n0) :: nil) = true -> match_block c tc pc (RBexit None (RBgoto n0) :: nil)
pc: node
m: memory_chunk
a0: Op.addressing
l: list reg
r: reg
n: RTL.node
Heqo: c ! pc = Some (RTL.Istore m a0 l r n)
m0: memory_chunk
a1: Op.addressing
l0: list reg
r0: reg
H: check_valid_node tc n0 && ceq peq n0 n && ceq memory_chunk_eq m0 m && ceq addressing_eq a1 a0 && ceq list_pos_eq l0 l && ceq peq r0 r = true

match_block c tc pc (RBstore None m0 a1 l0 r0 :: RBexit None (RBgoto n0) :: nil)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
n: RTL.node
IHb: forall pc : node, check_code c tc pc (RBexit None (RBgoto n) :: nil) = true -> match_block c tc pc (RBexit None (RBgoto n) :: nil)
pc: node
m: memory_chunk
a0: Op.addressing
l: list reg
r: reg
Heqo: c ! pc = Some (RTL.Istore m a0 l r n)
H0: check_valid_node tc n = true

match_block c tc pc (RBstore None m a0 l r :: RBexit None (RBgoto n) :: nil)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
n: RTL.node
IHb: forall pc : node, check_code c tc pc (RBexit None (RBgoto n) :: nil) = true -> match_block c tc pc (RBexit None (RBgoto n) :: nil)
pc: node
m: memory_chunk
a0: Op.addressing
l: list reg
r: reg
Heqo: c ! pc = Some (RTL.Istore m a0 l r n)
H0: check_valid_node tc n = true

valid_succ tc n
eauto using check_valid_node_correct.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
s: signature
s0: (reg + ident)%type
l: list reg
r: reg
n: RTL.node
Heqo: c ! pc = Some (RTL.Icall s s0 l r n)
H: match s0 with | inl r' => match a with | RBnop => match b with | RBexit None (RBcall sig (inl r0) args dst pc') :: nil => check_valid_node tc pc' && ceq peq pc' n && ceq signature_eq sig s && ceq peq r0 r' && ceq list_pos_eq args l && ceq peq dst r | RBexit None (RBcall sig (inl r0) args dst pc') :: _ :: _ => false | RBexit None (RBcall sig (inr _) _ _ _) :: _ => false | _ => false end | _ => false end | inr i' => match a with | RBnop => match b with | RBexit None (RBcall sig (inl _) _ _ _) :: _ => false | RBexit None (RBcall sig (inr i) args dst pc') :: nil => check_valid_node tc pc' && ceq peq pc' n && ceq signature_eq sig s && ceq peq i i' && ceq list_pos_eq args l && ceq peq dst r | RBexit None (RBcall sig (inr i) args dst pc') :: _ :: _ => false | _ => false end | _ => false end end = true

match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
s1: signature
l1: list reg
r1: reg
n0: node
r2: reg
IHb: forall pc : node, check_code c tc pc (RBexit None (RBcall s1 (inl r2) l1 r1 n0) :: nil) = true -> match_block c tc pc (RBexit None (RBcall s1 (inl r2) l1 r1 n0) :: nil)
pc: node
s: signature
l: list reg
r: reg
n: RTL.node
r0: reg
Heqo: c ! pc = Some (RTL.Icall s (inl r0) l r n)
H: check_valid_node tc n0 && ceq peq n0 n && ceq signature_eq s1 s && ceq peq r2 r0 && ceq list_pos_eq l1 l && ceq peq r1 r = true

match_block c tc pc (RBnop :: RBexit None (RBcall s1 (inl r2) l1 r1 n0) :: nil)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
s1: signature
l1: list reg
r0: reg
n0: node
i1: ident
IHb: forall pc : node, check_code c tc pc (RBexit None (RBcall s1 (inr i1) l1 r0 n0) :: nil) = true -> match_block c tc pc (RBexit None (RBcall s1 (inr i1) l1 r0 n0) :: nil)
pc: node
s: signature
l: list reg
r: reg
n: RTL.node
i: ident
Heqo: c ! pc = Some (RTL.Icall s (inr i) l r n)
H: check_valid_node tc n0 && ceq peq n0 n && ceq signature_eq s1 s && ceq peq i1 i && ceq list_pos_eq l1 l && ceq peq r0 r = true
match_block c tc pc (RBnop :: RBexit None (RBcall s1 (inr i1) l1 r0 n0) :: nil)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
n: RTL.node
s: signature
r0: reg
l: list reg
r: reg
IHb: forall pc : node, check_code c tc pc (RBexit None (RBcall s (inl r0) l r n) :: nil) = true -> match_block c tc pc (RBexit None (RBcall s (inl r0) l r n) :: nil)
pc: node
Heqo: c ! pc = Some (RTL.Icall s (inl r0) l r n)
H0: check_valid_node tc n = true

match_block c tc pc (RBnop :: RBexit None (RBcall s (inl r0) l r n) :: nil)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
s1: signature
l1: list reg
r0: reg
n0: node
i1: ident
IHb: forall pc : node, check_code c tc pc (RBexit None (RBcall s1 (inr i1) l1 r0 n0) :: nil) = true -> match_block c tc pc (RBexit None (RBcall s1 (inr i1) l1 r0 n0) :: nil)
pc: node
s: signature
l: list reg
r: reg
n: RTL.node
i: ident
Heqo: c ! pc = Some (RTL.Icall s (inr i) l r n)
H: check_valid_node tc n0 && ceq peq n0 n && ceq signature_eq s1 s && ceq peq i1 i && ceq list_pos_eq l1 l && ceq peq r0 r = true
match_block c tc pc (RBnop :: RBexit None (RBcall s1 (inr i1) l1 r0 n0) :: nil)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
s1: signature
l1: list reg
r0: reg
n0: node
i1: ident
IHb: forall pc : node, check_code c tc pc (RBexit None (RBcall s1 (inr i1) l1 r0 n0) :: nil) = true -> match_block c tc pc (RBexit None (RBcall s1 (inr i1) l1 r0 n0) :: nil)
pc: node
s: signature
l: list reg
r: reg
n: RTL.node
i: ident
Heqo: c ! pc = Some (RTL.Icall s (inr i) l r n)
H: check_valid_node tc n0 && ceq peq n0 n && ceq signature_eq s1 s && ceq peq i1 i && ceq list_pos_eq l1 l && ceq peq r0 r = true

match_block c tc pc (RBnop :: RBexit None (RBcall s1 (inr i1) l1 r0 n0) :: nil)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
n: RTL.node
s: signature
i: ident
l: list reg
r: reg
IHb: forall pc : node, check_code c tc pc (RBexit None (RBcall s (inr i) l r n) :: nil) = true -> match_block c tc pc (RBexit None (RBcall s (inr i) l r n) :: nil)
pc: node
Heqo: c ! pc = Some (RTL.Icall s (inr i) l r n)
H0: check_valid_node tc n = true

match_block c tc pc (RBnop :: RBexit None (RBcall s (inr i) l r n) :: nil)
eapply match_block_call; eauto using check_valid_node_correct.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
s: signature
s0: (reg + ident)%type
l: list reg
Heqo: c ! pc = Some (RTL.Itailcall s s0 l)
H: match s0 with | inl r => match a with | RBnop => match b with | RBexit None (RBtailcall sig' (inl r') args') :: nil => ceq signature_eq s sig' && ceq peq r r' && ceq list_pos_eq l args' | RBexit None (RBtailcall sig' (inl r') args') :: _ :: _ => false | RBexit None (RBtailcall sig' (inr _) _) :: _ => false | _ => false end | _ => false end | inr r => match a with | RBnop => match b with | RBexit None (RBtailcall sig' (inl _) _) :: _ => false | RBexit None (RBtailcall sig' (inr r') args') :: nil => ceq signature_eq s sig' && ceq peq r r' && ceq list_pos_eq l args' | RBexit None (RBtailcall sig' (inr r') args') :: _ :: _ => false | _ => false end | _ => false end end = true

match_block c tc pc (a :: b)
repeat (destruct_match; try discriminate); subst; try solve [unfold_ands; econstructor; eauto].
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
c0: Op.condition
l: list reg
n, n0: RTL.node
Heqo: c ! pc = Some (RTL.Icond c0 l n n0)
H: match a with | RBnop => match b with | RBexit None (RBcond cond' args' n1' n2') :: nil => check_valid_node tc n && check_valid_node tc n0 && ceq condition_eq c0 cond' && ceq list_pos_eq l args' && ceq peq n n1' && ceq peq n0 n2' | RBexit None (RBcond cond' args' n1' n2') :: _ :: _ => false | _ => false end | _ => false end = true

match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
c2: Op.condition
l1: list reg
n1, n2: node
IHb: forall pc : node, check_code c tc pc (RBexit None (RBcond c2 l1 n1 n2) :: nil) = true -> match_block c tc pc (RBexit None (RBcond c2 l1 n1 n2) :: nil)
pc: node
c0: Op.condition
l: list reg
n, n0: RTL.node
Heqo: c ! pc = Some (RTL.Icond c0 l n n0)
H: check_valid_node tc n && check_valid_node tc n0 && ceq condition_eq c0 c2 && ceq list_pos_eq l l1 && ceq peq n n1 && ceq peq n0 n2 = true

match_block c tc pc (RBnop :: RBexit None (RBcond c2 l1 n1 n2) :: nil)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
c2: Op.condition
l1: list reg
n1, n2: node
IHb: forall pc : node, check_code c tc pc (RBexit None (RBcond c2 l1 n1 n2) :: nil) = true -> match_block c tc pc (RBexit None (RBcond c2 l1 n1 n2) :: nil)
pc: node
Heqo: c ! pc = Some (RTL.Icond c2 l1 n1 n2)
H0: check_valid_node tc n1 = true
H5: check_valid_node tc n2 = true

match_block c tc pc (RBnop :: RBexit None (RBcond c2 l1 n1 n2) :: nil)
eapply match_block_cond; eauto using check_valid_node_correct.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
r: reg
l: list RTL.node
Heqo: c ! pc = Some (RTL.Ijumptable r l)
H: match a with | RBnop => match b with | RBexit None (RBjumptable r' ns') :: nil => ceq peq r r' && ceq list_pos_eq l ns' && forallb (check_valid_node tc) l | RBexit None (RBjumptable r' ns') :: _ :: _ => false | _ => false end | _ => false end = true

match_block c tc pc (a :: b)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
r0: reg
l1: list node
IHb: forall pc : node, check_code c tc pc (RBexit None (RBjumptable r0 l1) :: nil) = true -> match_block c tc pc (RBexit None (RBjumptable r0 l1) :: nil)
pc: node
r: reg
l: list RTL.node
Heqo: c ! pc = Some (RTL.Ijumptable r l)
H: ceq peq r r0 && ceq list_pos_eq l l1 && forallb (check_valid_node tc) l = true

match_block c tc pc (RBnop :: RBexit None (RBjumptable r0 l1) :: nil)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
r0: reg
l1: list node
IHb: forall pc : node, check_code c tc pc (RBexit None (RBjumptable r0 l1) :: nil) = true -> match_block c tc pc (RBexit None (RBjumptable r0 l1) :: nil)
pc: node
H1: forallb (check_valid_node tc) l1 = true
Heqo: c ! pc = Some (RTL.Ijumptable r0 l1)

match_block c tc pc (RBnop :: RBexit None (RBjumptable r0 l1) :: nil)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
r0: reg
l1: list node
IHb: forall pc : node, check_code c tc pc (RBexit None (RBjumptable r0 l1) :: nil) = true -> match_block c tc pc (RBexit None (RBjumptable r0 l1) :: nil)
pc: node
H1: forallb (check_valid_node tc) l1 = true
Heqo: c ! pc = Some (RTL.Ijumptable r0 l1)

Forall (valid_succ tc) l1
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
r0: reg
l1: list node
IHb: forall pc : node, check_code c tc pc (RBexit None (RBjumptable r0 l1) :: nil) = true -> match_block c tc pc (RBexit None (RBjumptable r0 l1) :: nil)
pc: node
H1: forallb (check_valid_node tc) l1 = true
Heqo: c ! pc = Some (RTL.Ijumptable r0 l1)
x: node
H: In x l1

valid_succ tc x
eapply forallb_forall in H1; eauto using check_valid_node_correct.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
c: RTL.code
tc: code
a: instr
b: list instr
IHb: forall pc : node, check_code c tc pc b = true -> match_block c tc pc b
pc: node
o: option reg
Heqo: c ! pc = Some (RTL.Ireturn o)
H: match o with | Some r => match a with | RBnop => match b with | RBexit None (RBreturn (Some r')) :: nil => ceq peq r r' | RBexit None (RBreturn (Some r')) :: _ :: _ => false | _ => false end | _ => false end | None => match a with | RBnop => match b with | RBexit None (RBreturn None) :: nil => true | _ => false end | _ => false end end = true

match_block c tc pc (a :: b)
repeat (destruct_match; try discriminate); subst; try solve [unfold_ands; econstructor; eauto]. Qed.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (f : RTL.function) (tf : function), transl_function f = OK tf -> match_code (RTL.fn_code f) (fn_code tf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (f : RTL.function) (tf : function), transl_function f = OK tf -> match_code (RTL.fn_code f) (fn_code tf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.function
tf: function
H: match partition f with | OK f' => if check_valid_node (fn_code f') (RTL.fn_entrypoint f) then if forall_ptree (check_code (RTL.fn_code f) (fn_code f')) (fn_code f') then OK {| fn_sig := RTL.fn_sig f; fn_params := RTL.fn_params f; fn_stacksize := RTL.fn_stacksize f; fn_code := fn_code f'; fn_entrypoint := RTL.fn_entrypoint f |} else Error (msg "check_present_blocks failed") else Error (msg "entrypoint does not exists") | Error msg => Error msg end = OK tf

match_code (RTL.fn_code f) (fn_code tf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.function
tf, f0: function
Heqr: partition f = OK f0
Heqb: check_valid_node (fn_code f0) (RTL.fn_entrypoint f) = true
Heqb0: forall_ptree (check_code (RTL.fn_code f) (fn_code f0)) (fn_code f0) = true
H: OK {| fn_sig := RTL.fn_sig f; fn_params := RTL.fn_params f; fn_stacksize := RTL.fn_stacksize f; fn_code := fn_code f0; fn_entrypoint := RTL.fn_entrypoint f |} = OK tf

match_code (RTL.fn_code f) (fn_code tf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.function
f0: function
Heqr: partition f = OK f0
Heqb: check_valid_node (fn_code f0) (RTL.fn_entrypoint f) = true
Heqb0: forall_ptree (check_code (RTL.fn_code f) (fn_code f0)) (fn_code f0) = true

match_code (RTL.fn_code f) (fn_code {| fn_sig := RTL.fn_sig f; fn_params := RTL.fn_params f; fn_stacksize := RTL.fn_stacksize f; fn_code := fn_code f0; fn_entrypoint := RTL.fn_entrypoint f |})
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.function
f0: function
Heqr: partition f = OK f0
Heqb: check_valid_node (fn_code f0) (RTL.fn_entrypoint f) = true
Heqb0: forall_ptree (check_code (RTL.fn_code f) (fn_code f0)) (fn_code f0) = true

match_code (RTL.fn_code f) (fn_code f0)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.function
f0: function
Heqr: partition f = OK f0
Heqb: check_valid_node (fn_code f0) (RTL.fn_entrypoint f) = true
Heqb0: forall_ptree (check_code (RTL.fn_code f) (fn_code f0)) (fn_code f0) = true
pc: positive
b: SeqBB.t
H: (fn_code f0) ! pc = Some b

match_block (RTL.fn_code f) (fn_code f0) pc b
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.function
f0: function
Heqr: partition f = OK f0
Heqb: check_valid_node (fn_code f0) (RTL.fn_entrypoint f) = true
pc: positive
b: SeqBB.t
H: (fn_code f0) ! pc = Some b
Heqb0: check_code (RTL.fn_code f) (fn_code f0) pc b = true

match_block (RTL.fn_code f) (fn_code f0) pc b
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
f: RTL.function
f0: function
Heqr: partition f = OK f0
Heqb: check_valid_node (fn_code f0) (RTL.fn_entrypoint f) = true
pc: positive
b: SeqBB.t
H: (fn_code f0) ! pc = Some b
Heqb0: (fix check_code (c : RTL.code) (tc : code) (pc : node) (b : SeqBB.t) {struct b} : bool := match c ! pc with | Some (RTL.Inop pc'') => match b with | RBnop :: (RBnop :: _ :: _) as b' | RBnop :: (RBop _ _ _ _ :: _ :: _) as b' | RBnop :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBnop :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBnop :: (RBexit (Some _) _ :: _ :: _) as b' | RBnop :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBnop :: (RBexit None (RBreturn _) :: _ :: _) as b' => check_code c tc pc'' b' | RBnop :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc'' | RBnop :: (RBexit None (RBgoto pc') :: _ :: _) as b' => check_code c tc pc'' b' | _ => false end | Some (RTL.Iop op' args' dst' pc'') => match b with | RBop None op args dst :: nil | RBop None op args dst :: RBnop :: nil | RBop None op args dst :: RBop _ _ _ _ :: nil | RBop None op args dst :: RBload _ _ _ _ _ :: nil | RBop None op args dst :: RBstore _ _ _ _ _ :: nil | RBop None op args dst :: RBsetpred _ _ _ _ :: nil | RBop None op args dst :: RBexit (Some _) _ :: nil | RBop None op args dst :: RBexit None (RBcall _ _ _ _ _) :: nil | RBop None op args dst :: RBexit None (RBtailcall _ _ _) :: nil | RBop None op args dst :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBop None op args dst :: RBexit None (RBcond _ _ _ _) :: nil | RBop None op args dst :: RBexit None (RBjumptable _ _) :: nil | RBop None op args dst :: RBexit None (RBreturn _) :: nil => false | RBop None op args dst :: (RBnop :: _ :: _) as b' | RBop None op args dst :: (RBop _ _ _ _ :: _ :: _) as b' | RBop None op args dst :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBop None op args dst :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBop None op args dst :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBop None op args dst :: (RBexit (Some _) _ :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBop None op args dst :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args args' && ceq peq dst dst' && check_code c tc pc'' b' | RBop None op args dst :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc'' && ceq operation_eq op op' && ceq list_pos_eq args args' && ceq peq dst dst' | RBop None op args dst :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq operation_eq op op' && ceq list_pos_eq args args' && ceq peq dst dst' && check_code c tc pc'' b' | _ => false end | Some (RTL.Iload chunk' addr' args' dst' pc'') => match b with | RBload None chunk addr args dst :: nil | RBload None chunk addr args dst :: RBnop :: nil | RBload None chunk addr args dst :: RBop _ _ _ _ :: nil | RBload None chunk addr args dst :: RBload _ _ _ _ _ :: nil | RBload None chunk addr args dst :: RBstore _ _ _ _ _ :: nil | RBload None chunk addr args dst :: RBsetpred _ _ _ _ :: nil | RBload None chunk addr args dst :: RBexit (Some _) _ :: nil | RBload None chunk addr args dst :: RBexit None (RBcall _ _ _ _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBtailcall _ _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBcond _ _ _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBjumptable _ _) :: nil | RBload None chunk addr args dst :: RBexit None (RBreturn _) :: nil => false | RBload None chunk addr args dst :: (RBnop :: _ :: _) as b' | RBload None chunk addr args dst :: (RBop _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit (Some _) _ :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBload None chunk addr args dst :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq dst dst' && check_code c tc pc'' b' | RBload None chunk addr args dst :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc'' && ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq dst dst' | RBload None chunk addr args dst :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq dst dst' && check_code c tc pc'' b' | _ => false end | Some (RTL.Istore chunk' addr' args' src' pc'') => match b with | RBstore None chunk addr args src :: nil | RBstore None chunk addr args src :: RBnop :: nil | RBstore None chunk addr args src :: RBop _ _ _ _ :: nil | RBstore None chunk addr args src :: RBload _ _ _ _ _ :: nil | RBstore None chunk addr args src :: RBstore _ _ _ _ _ :: nil | RBstore None chunk addr args src :: RBsetpred _ _ _ _ :: nil | RBstore None chunk addr args src :: RBexit (Some _) _ :: nil | RBstore None chunk addr args src :: RBexit None (RBcall _ _ _ _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBtailcall _ _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBbuiltin _ _ _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBcond _ _ _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBjumptable _ _) :: nil | RBstore None chunk addr args src :: RBexit None (RBreturn _) :: nil => false | RBstore None chunk addr args src :: (RBnop :: _ :: _) as b' | RBstore None chunk addr args src :: (RBop _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBload _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBstore _ _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBsetpred _ _ _ _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit (Some _) _ :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBcall _ _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBtailcall _ _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBbuiltin _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBcond _ _ _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBjumptable _ _) :: _ :: _) as b' | RBstore None chunk addr args src :: (RBexit None (RBreturn _) :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq src src' && check_code c tc pc'' b' | RBstore None chunk addr args src :: RBexit None (RBgoto pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc'' && ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq src src' | RBstore None chunk addr args src :: (RBexit None (RBgoto pc') :: _ :: _) as b' => ceq memory_chunk_eq chunk chunk' && ceq addressing_eq addr addr' && ceq list_pos_eq args args' && ceq peq src src' && check_code c tc pc'' b' | _ => false end | Some (RTL.Icall sig' (inl r') args' dst' pc'') => match b with | RBnop :: RBexit None (RBcall sig (inl r) args dst pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc'' && ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args args' && ceq peq dst dst' | RBnop :: RBexit None (RBcall sig (inl r) args dst pc') :: _ :: _ => false | RBnop :: RBexit None (RBcall sig (inr _) _ _ _) :: _ => false | _ => false end | Some (RTL.Icall sig' (inr i') args' dst' pc'') => match b with | RBnop :: RBexit None (RBcall sig (inl _) _ _ _) :: _ => false | RBnop :: RBexit None (RBcall sig (inr i) args dst pc') :: nil => check_valid_node tc pc' && ceq peq pc' pc'' && ceq signature_eq sig sig' && ceq peq i i' && ceq list_pos_eq args args' && ceq peq dst dst' | RBnop :: RBexit None (RBcall sig (inr i) args dst pc') :: _ :: _ => false | _ => false end | Some (RTL.Itailcall sig (inl r) args) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl r') args') :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args args' | RBnop :: RBexit None (RBtailcall sig' (inl r') args') :: _ :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr _) _) :: _ => false | _ => false end | Some (RTL.Itailcall sig (inr r) args) => match b with | RBnop :: RBexit None (RBtailcall sig' (inl _) _) :: _ => false | RBnop :: RBexit None (RBtailcall sig' (inr r') args') :: nil => ceq signature_eq sig sig' && ceq peq r r' && ceq list_pos_eq args args' | RBnop :: RBexit None (RBtailcall sig' (inr r') args') :: _ :: _ => false | _ => false end | Some (RTL.Icond cond args n1 n2) => match b with | RBnop :: RBexit None (RBcond cond' args' n1' n2') :: nil => check_valid_node tc n1 && check_valid_node tc n2 && ceq condition_eq cond cond' && ceq list_pos_eq args args' && ceq peq n1 n1' && ceq peq n2 n2' | RBnop :: RBexit None (RBcond cond' args' n1' n2') :: _ :: _ => false | _ => false end | Some (RTL.Ijumptable r ns) => match b with | RBnop :: RBexit None (RBjumptable r' ns') :: nil => ceq peq r r' && ceq list_pos_eq ns ns' && forallb (check_valid_node tc) ns | RBnop :: RBexit None (RBjumptable r' ns') :: _ :: _ => false | _ => false end | Some (RTL.Ireturn (Some r)) => match b with | RBnop :: RBexit None (RBreturn (Some r')) :: nil => ceq peq r r' | RBnop :: RBexit None (RBreturn (Some r')) :: _ :: _ => false | _ => false end | Some (RTL.Ireturn None) => match b with | RBnop :: RBexit None (RBreturn None) :: nil => true | _ => false end | _ => false end) (RTL.fn_code f) (fn_code f0) pc b = true

match_block (RTL.fn_code f) (fn_code f0) pc b
eauto using transl_match_code'. Qed.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall s1 : RTL.state, RTL.initial_state prog s1 -> exists s2 : state, initial_state tprog s2 /\ match_states None s1 s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall s1 : RTL.state, RTL.initial_state prog s1 -> exists s2 : state, initial_state tprog s2 /\ match_states None s1 s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main

exists s2 : state, initial_state tprog s2 /\ match_states None (RTL.Callstate nil f nil m0) s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main

(exists tf : fundef, Genv.find_funct_ptr tge b = Some tf /\ transl_fundef f = OK tf) -> exists s2 : state, initial_state tprog s2 /\ match_states None (RTL.Callstate nil f nil m0) s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf

exists s2 : state, initial_state tprog s2 /\ match_states None (RTL.Callstate nil f nil m0) s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf

initial_state tprog ?s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
match_states None (RTL.Callstate nil f nil m0) ?s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf

initial_state tprog ?s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
match_states None (RTL.Callstate nil f nil m0) ?s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf

Genv.init_mem tprog = Some ?m0
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
Genv.find_symbol (Genv.globalenv tprog) (prog_main tprog) = Some ?b
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
Genv.find_funct_ptr (Genv.globalenv tprog) ?b = Some ?f
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
funsig ?f = signature_main
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
match_states None (RTL.Callstate nil f nil m0) (Callstate nil ?f nil ?m0)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf

Genv.find_symbol (Genv.globalenv tprog) (prog_main tprog) = Some ?b
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
Genv.find_funct_ptr (Genv.globalenv tprog) ?b = Some ?f
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
funsig ?f = signature_main
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
match_states None (RTL.Callstate nil f nil m0) (Callstate nil ?f nil m0)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf

Genv.find_symbol (Genv.globalenv tprog) (prog_main prog) = Some ?b
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
prog_main prog = prog_main tprog
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
Genv.find_funct_ptr (Genv.globalenv tprog) ?b = Some ?f
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
funsig ?f = signature_main
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
match_states None (RTL.Callstate nil f nil m0) (Callstate nil ?f nil m0)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf

prog_main prog = prog_main tprog
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
Genv.find_funct_ptr (Genv.globalenv tprog) b = Some ?f
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
funsig ?f = signature_main
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
match_states None (RTL.Callstate nil f nil m0) (Callstate nil ?f nil m0)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf

Genv.find_funct_ptr (Genv.globalenv tprog) b = Some ?f
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
funsig ?f = signature_main
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
match_states None (RTL.Callstate nil f nil m0) (Callstate nil ?f nil m0)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf

funsig tf = signature_main
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
match_states None (RTL.Callstate nil f nil m0) (Callstate nil tf nil m0)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf

match_states None (RTL.Callstate nil f nil m0) (Callstate nil tf nil m0)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf

transl_fundef f = OK tf
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf
Forall2 match_stackframe nil nil
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
b: block
f: RTL.fundef
m0: mem
ge0:= Genv.globalenv prog: Genv.t RTL.fundef unit
H: Genv.init_mem prog = Some m0
H0: Genv.find_symbol ge0 (prog_main prog) = Some b
H1: Genv.find_funct_ptr ge0 b = Some f
H2: RTL.funsig f = signature_main
tf: fundef
A: Genv.find_funct_ptr tge b = Some tf
B: transl_fundef f = OK tf

Forall2 match_stackframe nil nil
auto. Qed.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (s1 : RTL.state) (s2 : state) (r : Integers.Int.int) (b : option SeqBB.t), match_states b s1 s2 -> RTL.final_state s1 r -> final_state s2 r
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (s1 : RTL.state) (s2 : state) (r : Integers.Int.int) (b : option SeqBB.t), match_states b s1 s2 -> RTL.final_state s1 r -> final_state s2 r
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s1: RTL.state
s2: state
r: Integers.Int.int
b: option SeqBB.t
H: match_states b s1 s2
H0: RTL.final_state s1 r
r0: Integers.Int.int
m: mem
H1: RTL.Returnstate nil (Vint r0) m = s1
H2: r0 = r

final_state s2 r
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s2: state
r: Integers.Int.int
b: option SeqBB.t
m: mem
H: match_states b (RTL.Returnstate nil (Vint r) m) s2
H0: RTL.final_state (RTL.Returnstate nil (Vint r) m) r

final_state s2 r
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
r: Integers.Int.int
m: mem
H0: RTL.final_state (RTL.Returnstate nil (Vint r) m) r
cs': list stackframe
STK: Forall2 match_stackframe nil cs'

final_state (Returnstate cs' (Vint r) m) r
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
r: Integers.Int.int
m: mem
H0: RTL.final_state (RTL.Returnstate nil (Vint r) m) r

final_state (Returnstate nil (Vint r) m) r
econstructor. Qed.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (A : Type) (n : nat) (l : list A), hd_error (list_drop n l) = nth_error l n
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (A : Type) (n : nat) (l : list A), hd_error (list_drop n l) = nth_error l n
induction n; destruct l; crush. Qed.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (A : Type) (l : list A) (n : A), hd_error l = Some n -> l = n :: tl l
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (A : Type) (l : list A) (n : A), hd_error l = Some n -> l = n :: tl l
induction l; crush. Qed.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
pc, pc': node

{imm_succ pc pc'} + {~ imm_succ pc pc'}
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
pc, pc': node

{imm_succ pc pc'} + {~ imm_succ pc pc'}
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
pc, pc': node

{pc' = Pos.pred pc} + {pc' <> Pos.pred pc}
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
pc, pc': node
H: forall x y : positive, {x = y} + {x <> y}

{pc' = Pos.pred pc} + {pc' <> Pos.pred pc}
decide equality. Defined.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
pc: node

imm_succ pc (Pos.pred pc)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
pc: node

imm_succ pc (Pos.pred pc)
unfold imm_succ; auto. Qed.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
pc: positive

imm_succ (Pos.succ pc) pc
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
pc: positive

imm_succ (Pos.succ pc) pc
unfold imm_succ; auto using Pos.pred_succ. Qed.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (s1 : state) (b : option SeqBB.t) (t : trace) (s : RTL.state), (exists (b' : option SeqBB.t) (s2 : state), star step tge s1 t s2 /\ ltof (option SeqBB.t) measure b' b /\ match_states b' s s2) -> exists (b' : option SeqBB.t) (s2 : state), (plus step tge s1 t s2 \/ star step tge s1 t s2 /\ ltof (option SeqBB.t) measure b' b) /\ match_states b' s s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (s1 : state) (b : option SeqBB.t) (t : trace) (s : RTL.state), (exists (b' : option SeqBB.t) (s2 : state), star step tge s1 t s2 /\ ltof (option SeqBB.t) measure b' b /\ match_states b' s s2) -> exists (b' : option SeqBB.t) (s2 : state), (plus step tge s1 t s2 \/ star step tge s1 t s2 /\ ltof (option SeqBB.t) measure b' b) /\ match_states b' s s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s1: state
b: option SeqBB.t
t: trace
s: RTL.state
x: option SeqBB.t
x0: state
H0: star step tge s1 t x0
H: ltof (option SeqBB.t) measure x b
H2: match_states x s x0

exists (b' : option SeqBB.t) (s2 : state), (plus step tge s1 t s2 \/ star step tge s1 t s2 /\ ltof (option SeqBB.t) measure b' b) /\ match_states b' s s2
do 3 econstructor; eauto. Qed.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (s1 : state) (b : option SeqBB.t) (t : trace) (s : RTL.state), (exists (b' : option SeqBB.t) (s2 : state), plus step tge s1 t s2 /\ match_states b' s s2) -> exists (b' : option SeqBB.t) (s2 : state), (plus step tge s1 t s2 \/ star step tge s1 t s2 /\ ltof (option SeqBB.t) measure b' b) /\ match_states b' s s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (s1 : state) (b : option SeqBB.t) (t : trace) (s : RTL.state), (exists (b' : option SeqBB.t) (s2 : state), plus step tge s1 t s2 /\ match_states b' s s2) -> exists (b' : option SeqBB.t) (s2 : state), (plus step tge s1 t s2 \/ star step tge s1 t s2 /\ ltof (option SeqBB.t) measure b' b) /\ match_states b' s s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s1: state
b: option SeqBB.t
t: trace
s: RTL.state
x: option SeqBB.t
x0: state
H0: plus step tge s1 t x0
H1: match_states x s x0

exists (b' : option SeqBB.t) (s2 : state), (plus step tge s1 t s2 \/ star step tge s1 t s2 /\ ltof (option SeqBB.t) measure b' b) /\ match_states b' s s2
do 3 econstructor; eauto. Qed.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (s : list RTL.stackframe) (f : RTL.function) (sp : val) (pc : positive) (rs : RTL.regset) (m : mem) (pc' : RTL.node), (RTL.fn_code f) ! pc = Some (RTL.Inop pc') -> forall (b : option SeqBB.t) (s2 : state), match_states b (RTL.State s f sp pc rs m) s2 -> exists (b' : option SeqBB.t) (s2' : state), (plus step tge s2 E0 s2' \/ star step tge s2 E0 s2' /\ ltof (option SeqBB.t) measure b' b) /\ match_states b' (RTL.State s f sp pc' rs m) s2'
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog

forall (s : list RTL.stackframe) (f : RTL.function) (sp : val) (pc : positive) (rs : RTL.regset) (m : mem) (pc' : RTL.node), (RTL.fn_code f) ! pc = Some (RTL.Inop pc') -> forall (b : option SeqBB.t) (s2 : state), match_states b (RTL.State s f sp pc rs m) s2 -> exists (b' : option SeqBB.t) (s2' : state), (plus step tge s2 E0 s2' \/ star step tge s2 E0 s2' /\ ltof (option SeqBB.t) measure b' b) /\ match_states b' (RTL.State s f sp pc' rs m) s2'
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')

forall (b : option SeqBB.t) (s2 : state), match_states b (RTL.State s f sp pc rs m) s2 -> exists (b' : option SeqBB.t) (s2' : state), (plus step tge s2 E0 s2' \/ star step tge s2 E0 s2' /\ ltof (option SeqBB.t) measure b' b) /\ match_states b' (RTL.State s f sp pc' rs m) s2'
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
b0: SeqBB.t
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
H0: match_states (Some b0) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: match_code (RTL.fn_code f) (fn_code tf)
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) b0
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H3: match_block (RTL.fn_code f) (fn_code tf) pc b0

exists (b' : option SeqBB.t) (s2' : state), (plus step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2' \/ star step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2' /\ ltof (option SeqBB.t) measure b' (Some b0)) /\ match_states b' (RTL.State s f sp pc' rs m) s2'
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
b0: SeqBB.t
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
H0: match_states (Some b0) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) b0
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H3: match_block (RTL.fn_code f) (fn_code tf) pc b0

exists (b' : option SeqBB.t) (s2' : state), (plus step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2' \/ star step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2' /\ ltof (option SeqBB.t) measure b' (Some b0)) /\ match_states b' (RTL.State s f sp pc' rs m) s2'
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H

exists (b' : option SeqBB.t) (s2' : state), (plus step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2' \/ star step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2' /\ ltof (option SeqBB.t) measure b' (Some (RBnop :: b))) /\ match_states b' (RTL.State s f sp pc' rs m) s2'
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
H0: match_states (Some (RBnop :: RBexit None (RBgoto pc') :: nil)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: valid_succ (fn_code tf) pc'
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
exists (b' : option SeqBB.t) (s2' : state), (plus step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2' \/ star step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2' /\ ltof (option SeqBB.t) measure b' (Some (RBnop :: RBexit None (RBgoto pc') :: nil))) /\ match_states b' (RTL.State s f sp pc' rs m) s2'
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H

exists (b' : option SeqBB.t) (s2' : state), (plus step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2' \/ star step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2' /\ ltof (option SeqBB.t) measure b' (Some (RBnop :: b))) /\ match_states b' (RTL.State s f sp pc' rs m) s2'
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H

exists (b' : option SeqBB.t) (s2 : state), star step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2 /\ ltof (option SeqBB.t) measure b' (Some (RBnop :: b)) /\ match_states b' (RTL.State s f sp pc' rs m) s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H

star step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 ?s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
ltof (option SeqBB.t) measure ?b' (Some (RBnop :: b)) /\ match_states ?b' (RTL.State s f sp pc' rs m) ?s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H

ltof (option SeqBB.t) measure ?b' (Some (RBnop :: b)) /\ match_states ?b' (RTL.State s f sp pc' rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H

ltof (option SeqBB.t) measure ?b' (Some (RBnop :: b))
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
match_states ?b' (RTL.State s f sp pc' rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H

match_states (Some b) (RTL.State s f sp pc' rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H

exists b' : SeqBB.t, (fn_code tf) ! pc1 = Some b' /\ match_block (RTL.fn_code f) (fn_code tf) pc' b
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc' rs m)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) b
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H

star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc' rs m)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) b
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H

RTL.step ge (RTL.State s f sp pc rs m) ?t2 (RTL.State s f sp pc' rs m)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
E0 = E0 ** ?t2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) b
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H

E0 = E0 ** E0
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) b
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: b)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H

sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) b
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: forall (out_s : istate) (block2 : SeqBB.t), SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (RBnop :: b) out_s -> (fn_code tf) ! pc1 = Some block2 -> SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H

forall (out_s : istate) (block2 : SeqBB.t), SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) b out_s -> (fn_code tf) ! pc1 = Some block2 -> SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
b: SeqBB.t
H0: match_states (Some (RBnop :: b)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: forall (out_s : istate) (block2 : SeqBB.t), SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (RBnop :: b) out_s -> (fn_code tf) ! pc1 = Some block2 -> SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' b
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
out_s: istate
block2: SeqBB.t
H3: SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) b out_s
H5: (fn_code tf) ! pc1 = Some block2

SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
i: instr
instrs: list instr
H0: match_states (Some (RBnop :: i :: instrs)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: forall (out_s : istate) (block2 : SeqBB.t), SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (RBnop :: i :: instrs) out_s -> (fn_code tf) ! pc1 = Some block2 -> SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' (i :: instrs)
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
block2: SeqBB.t
H5: (fn_code tf) ! pc1 = Some block2
state', state'': instr_state
cf: cf_instr
H7: step_instr tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) i (Iexec state')
H9: step_list (step_instr tge) sp (Iexec state') instrs (Iterm state'' cf)

SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 (Iterm state'' cf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
i: instr
instrs: list instr
H0: match_states (Some (RBnop :: i :: instrs)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: forall (out_s : istate) (block2 : SeqBB.t), SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (RBnop :: i :: instrs) out_s -> (fn_code tf) ! pc1 = Some block2 -> SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' (i :: instrs)
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
block2: SeqBB.t
H5: (fn_code tf) ! pc1 = Some block2
state': instr_state
cf: cf_instr
H8: step_instr tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) i (Iterm state' cf)
SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 (Iterm state' cf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
i: instr
instrs: list instr
H0: match_states (Some (RBnop :: i :: instrs)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: forall (out_s : istate) (block2 : SeqBB.t), SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (RBnop :: i :: instrs) out_s -> (fn_code tf) ! pc1 = Some block2 -> SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' (i :: instrs)
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
block2: SeqBB.t
H5: (fn_code tf) ! pc1 = Some block2
state', state'': instr_state
cf: cf_instr
H7: step_instr tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) i (Iexec state')
H9: step_list (step_instr tge) sp (Iexec state') instrs (Iterm state'' cf)

SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (RBnop :: i :: instrs) (Iterm state'' cf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
i: instr
instrs: list instr
H0: match_states (Some (RBnop :: i :: instrs)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: forall (out_s : istate) (block2 : SeqBB.t), SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (RBnop :: i :: instrs) out_s -> (fn_code tf) ! pc1 = Some block2 -> SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' (i :: instrs)
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
block2: SeqBB.t
H5: (fn_code tf) ! pc1 = Some block2
state', state'': instr_state
cf: cf_instr
H7: step_instr tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) i (Iexec state')
H9: step_list (step_instr tge) sp (Iexec state') instrs (Iterm state'' cf)
(fn_code tf) ! pc1 = Some block2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
i: instr
instrs: list instr
H0: match_states (Some (RBnop :: i :: instrs)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: forall (out_s : istate) (block2 : SeqBB.t), SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (RBnop :: i :: instrs) out_s -> (fn_code tf) ! pc1 = Some block2 -> SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' (i :: instrs)
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
block2: SeqBB.t
H5: (fn_code tf) ! pc1 = Some block2
state': instr_state
cf: cf_instr
H8: step_instr tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) i (Iterm state' cf)
SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 (Iterm state' cf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
i: instr
instrs: list instr
H0: match_states (Some (RBnop :: i :: instrs)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: forall (out_s : istate) (block2 : SeqBB.t), SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (RBnop :: i :: instrs) out_s -> (fn_code tf) ! pc1 = Some block2 -> SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' (i :: instrs)
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
block2: SeqBB.t
H5: (fn_code tf) ! pc1 = Some block2
state', state'': instr_state
cf: cf_instr
H7: step_instr tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) i (Iexec state')
H9: step_list (step_instr tge) sp (Iexec state') instrs (Iterm state'' cf)

step_instr tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) RBnop (Iexec ?state')
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
i: instr
instrs: list instr
H0: match_states (Some (RBnop :: i :: instrs)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: forall (out_s : istate) (block2 : SeqBB.t), SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (RBnop :: i :: instrs) out_s -> (fn_code tf) ! pc1 = Some block2 -> SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' (i :: instrs)
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
block2: SeqBB.t
H5: (fn_code tf) ! pc1 = Some block2
state', state'': instr_state
cf: cf_instr
H7: step_instr tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) i (Iexec state')
H9: step_list (step_instr tge) sp (Iexec state') instrs (Iterm state'' cf)
step_list (step_instr tge) sp (Iexec ?state') (i :: instrs) (Iterm state'' cf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
i: instr
instrs: list instr
H0: match_states (Some (RBnop :: i :: instrs)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: forall (out_s : istate) (block2 : SeqBB.t), SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (RBnop :: i :: instrs) out_s -> (fn_code tf) ! pc1 = Some block2 -> SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' (i :: instrs)
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
block2: SeqBB.t
H5: (fn_code tf) ! pc1 = Some block2
state', state'': instr_state
cf: cf_instr
H7: step_instr tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) i (Iexec state')
H9: step_list (step_instr tge) sp (Iexec state') instrs (Iterm state'' cf)
(fn_code tf) ! pc1 = Some block2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
i: instr
instrs: list instr
H0: match_states (Some (RBnop :: i :: instrs)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: forall (out_s : istate) (block2 : SeqBB.t), SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (RBnop :: i :: instrs) out_s -> (fn_code tf) ! pc1 = Some block2 -> SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' (i :: instrs)
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
block2: SeqBB.t
H5: (fn_code tf) ! pc1 = Some block2
state': instr_state
cf: cf_instr
H8: step_instr tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) i (Iterm state' cf)
SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 (Iterm state' cf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
i: instr
instrs: list instr
H0: match_states (Some (RBnop :: i :: instrs)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: forall (out_s : istate) (block2 : SeqBB.t), SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (RBnop :: i :: instrs) out_s -> (fn_code tf) ! pc1 = Some block2 -> SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' (i :: instrs)
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
block2: SeqBB.t
H5: (fn_code tf) ! pc1 = Some block2
state', state'': instr_state
cf: cf_instr
H7: step_instr tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) i (Iexec state')
H9: step_list (step_instr tge) sp (Iexec state') instrs (Iterm state'' cf)

step_list (step_instr tge) sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (i :: instrs) (Iterm state'' cf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
i: instr
instrs: list instr
H0: match_states (Some (RBnop :: i :: instrs)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: forall (out_s : istate) (block2 : SeqBB.t), SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (RBnop :: i :: instrs) out_s -> (fn_code tf) ! pc1 = Some block2 -> SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' (i :: instrs)
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
block2: SeqBB.t
H5: (fn_code tf) ! pc1 = Some block2
state', state'': instr_state
cf: cf_instr
H7: step_instr tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) i (Iexec state')
H9: step_list (step_instr tge) sp (Iexec state') instrs (Iterm state'' cf)
(fn_code tf) ! pc1 = Some block2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
i: instr
instrs: list instr
H0: match_states (Some (RBnop :: i :: instrs)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: forall (out_s : istate) (block2 : SeqBB.t), SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (RBnop :: i :: instrs) out_s -> (fn_code tf) ! pc1 = Some block2 -> SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' (i :: instrs)
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
block2: SeqBB.t
H5: (fn_code tf) ! pc1 = Some block2
state': instr_state
cf: cf_instr
H8: step_instr tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) i (Iterm state' cf)
SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 (Iterm state' cf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
i: instr
instrs: list instr
H0: match_states (Some (RBnop :: i :: instrs)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: forall (out_s : istate) (block2 : SeqBB.t), SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (RBnop :: i :: instrs) out_s -> (fn_code tf) ! pc1 = Some block2 -> SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' (i :: instrs)
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
block2: SeqBB.t
H5: (fn_code tf) ! pc1 = Some block2
state', state'': instr_state
cf: cf_instr
H7: step_instr tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) i (Iexec state')
H9: step_list (step_instr tge) sp (Iexec state') instrs (Iterm state'' cf)

(fn_code tf) ! pc1 = Some block2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
i: instr
instrs: list instr
H0: match_states (Some (RBnop :: i :: instrs)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: forall (out_s : istate) (block2 : SeqBB.t), SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (RBnop :: i :: instrs) out_s -> (fn_code tf) ! pc1 = Some block2 -> SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' (i :: instrs)
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
block2: SeqBB.t
H5: (fn_code tf) ! pc1 = Some block2
state': instr_state
cf: cf_instr
H8: step_instr tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) i (Iterm state' cf)
SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 (Iterm state' cf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
i: instr
instrs: list instr
H0: match_states (Some (RBnop :: i :: instrs)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: forall (out_s : istate) (block2 : SeqBB.t), SeqBB.step tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (RBnop :: i :: instrs) out_s -> (fn_code tf) ! pc1 = Some block2 -> SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 out_s
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: match_block (RTL.fn_code f) (fn_code tf) pc' (i :: instrs)
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
block2: SeqBB.t
H5: (fn_code tf) ! pc1 = Some block2
state': instr_state
cf: cf_instr
H8: step_instr tge sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) i (Iterm state' cf)

SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) block2 (Iterm state' cf)
inv H4; inv H8.
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
H0: match_states (Some (RBnop :: RBexit None (RBgoto pc') :: nil)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: valid_succ (fn_code tf) pc'
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H

exists (b' : option SeqBB.t) (s2' : state), (plus step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2' \/ star step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2' /\ ltof (option SeqBB.t) measure b' (Some (RBnop :: RBexit None (RBgoto pc') :: nil))) /\ match_states b' (RTL.State s f sp pc' rs m) s2'
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
H0: match_states (Some (RBnop :: RBexit None (RBgoto pc') :: nil)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: valid_succ (fn_code tf) pc'
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H

exists (b' : option SeqBB.t) (s2' : state), (plus step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2' \/ star step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2' /\ ltof (option SeqBB.t) measure b' (Some (RBnop :: RBexit None (RBgoto pc') :: nil))) /\ match_states b' (RTL.State s f sp pc' rs m) s2'
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
H0: match_states (Some (RBnop :: RBexit None (RBgoto pc') :: nil)) (RTL.State s f sp pc rs m) (State stk' tf sp pc1 rs1 (PMap.init false) m1)
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: valid_succ (fn_code tf) pc'
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H

exists (b' : option SeqBB.t) (s2 : state), plus step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2 /\ match_states b' (RTL.State s f sp pc' rs m) s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: valid_succ (fn_code tf) pc'
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
TF0: transl_function f = OK tf
CODE0: match_code (RTL.fn_code f) (fn_code tf)
BLOCK: exists b' : SeqBB.t, (fn_code tf) ! pc1 = Some b' /\ match_block (RTL.fn_code f) (fn_code tf) pc (RBnop :: RBexit None (RBgoto pc') :: nil)
STK0: Forall2 match_stackframe s stk'
STARSIMU0: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB0: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)

exists (b' : option SeqBB.t) (s2 : state), plus step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2 /\ match_states b' (RTL.State s f sp pc' rs m) s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H4: valid_succ (fn_code tf) pc'
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
TF0: transl_function f = OK tf
CODE0: match_code (RTL.fn_code f) (fn_code tf)
STK0: Forall2 match_stackframe s stk'
STARSIMU0: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB0: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
H3: (fn_code tf) ! pc1 = Some x
H5: match_block (RTL.fn_code f) (fn_code tf) pc (RBnop :: RBexit None (RBgoto pc') :: nil)
Learn0: Learnt H2

exists (b' : option SeqBB.t) (s2 : state), plus step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2 /\ match_states b' (RTL.State s f sp pc' rs m) s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
TF0: transl_function f = OK tf
CODE0: match_code (RTL.fn_code f) (fn_code tf)
STK0: Forall2 match_stackframe s stk'
STARSIMU0: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB0: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
H3: (fn_code tf) ! pc1 = Some x
H5: match_block (RTL.fn_code f) (fn_code tf) pc (RBnop :: RBexit None (RBgoto pc') :: nil)
Learn0: Learnt H2
x0: SeqBB.t
H0: (fn_code tf) ! pc' = Some x0

exists (b' : option SeqBB.t) (s2 : state), plus step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 s2 /\ match_states b' (RTL.State s f sp pc' rs m) s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
TF0: transl_function f = OK tf
CODE0: match_code (RTL.fn_code f) (fn_code tf)
STK0: Forall2 match_stackframe s stk'
STARSIMU0: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB0: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
H3: (fn_code tf) ! pc1 = Some x
H5: match_block (RTL.fn_code f) (fn_code tf) pc (RBnop :: RBexit None (RBgoto pc') :: nil)
Learn0: Learnt H2
x0: SeqBB.t
H0: (fn_code tf) ! pc' = Some x0

plus step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 ?s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
TF0: transl_function f = OK tf
CODE0: match_code (RTL.fn_code f) (fn_code tf)
STK0: Forall2 match_stackframe s stk'
STARSIMU0: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB0: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
H3: (fn_code tf) ! pc1 = Some x
H5: match_block (RTL.fn_code f) (fn_code tf) pc (RBnop :: RBexit None (RBgoto pc') :: nil)
Learn0: Learnt H2
x0: SeqBB.t
H0: (fn_code tf) ! pc' = Some x0
match_states ?b' (RTL.State s f sp pc' rs m) ?s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
TF0: transl_function f = OK tf
CODE0: match_code (RTL.fn_code f) (fn_code tf)
STK0: Forall2 match_stackframe s stk'
STARSIMU0: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB0: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
H3: (fn_code tf) ! pc1 = Some x
H5: match_block (RTL.fn_code f) (fn_code tf) pc (RBnop :: RBexit None (RBgoto pc') :: nil)
Learn0: Learnt H2
x0: SeqBB.t
H0: (fn_code tf) ! pc' = Some x0

step tge (State stk' tf sp pc1 rs1 (PMap.init false) m1) E0 ?s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
TF0: transl_function f = OK tf
CODE0: match_code (RTL.fn_code f) (fn_code tf)
STK0: Forall2 match_stackframe s stk'
STARSIMU0: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB0: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
H3: (fn_code tf) ! pc1 = Some x
H5: match_block (RTL.fn_code f) (fn_code tf) pc (RBnop :: RBexit None (RBgoto pc') :: nil)
Learn0: Learnt H2
x0: SeqBB.t
H0: (fn_code tf) ! pc' = Some x0
match_states ?b' (RTL.State s f sp pc' rs m) ?s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
TF0: transl_function f = OK tf
CODE0: match_code (RTL.fn_code f) (fn_code tf)
STK0: Forall2 match_stackframe s stk'
STARSIMU0: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB0: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
H3: (fn_code tf) ! pc1 = Some x
H5: match_block (RTL.fn_code f) (fn_code tf) pc (RBnop :: RBexit None (RBgoto pc') :: nil)
Learn0: Learnt H2
x0: SeqBB.t
H0: (fn_code tf) ! pc' = Some x0

(fn_code tf) ! pc1 = Some ?bb
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
TF0: transl_function f = OK tf
CODE0: match_code (RTL.fn_code f) (fn_code tf)
STK0: Forall2 match_stackframe s stk'
STARSIMU0: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB0: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
H3: (fn_code tf) ! pc1 = Some x
H5: match_block (RTL.fn_code f) (fn_code tf) pc (RBnop :: RBexit None (RBgoto pc') :: nil)
Learn0: Learnt H2
x0: SeqBB.t
H0: (fn_code tf) ! pc' = Some x0
SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) ?bb (Iterm {| is_rs := ?rs'; is_ps := ?pr'; is_mem := ?m' |} ?cf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
TF0: transl_function f = OK tf
CODE0: match_code (RTL.fn_code f) (fn_code tf)
STK0: Forall2 match_stackframe s stk'
STARSIMU0: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB0: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
H3: (fn_code tf) ! pc1 = Some x
H5: match_block (RTL.fn_code f) (fn_code tf) pc (RBnop :: RBexit None (RBgoto pc') :: nil)
Learn0: Learnt H2
x0: SeqBB.t
H0: (fn_code tf) ! pc' = Some x0
step_cf_instr tge (State stk' tf sp pc1 ?rs' ?pr' ?m') ?cf E0 ?s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
TF0: transl_function f = OK tf
CODE0: match_code (RTL.fn_code f) (fn_code tf)
STK0: Forall2 match_stackframe s stk'
STARSIMU0: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB0: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
H3: (fn_code tf) ! pc1 = Some x
H5: match_block (RTL.fn_code f) (fn_code tf) pc (RBnop :: RBexit None (RBgoto pc') :: nil)
Learn0: Learnt H2
x0: SeqBB.t
H0: (fn_code tf) ! pc' = Some x0
match_states ?b' (RTL.State s f sp pc' rs m) ?s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
TF0: transl_function f = OK tf
CODE0: match_code (RTL.fn_code f) (fn_code tf)
STK0: Forall2 match_stackframe s stk'
STARSIMU0: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB0: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
H3: (fn_code tf) ! pc1 = Some x
H5: match_block (RTL.fn_code f) (fn_code tf) pc (RBnop :: RBexit None (RBgoto pc') :: nil)
Learn0: Learnt H2
x0: SeqBB.t
H0: (fn_code tf) ! pc' = Some x0

SeqBB.step tge sp (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) x (Iterm {| is_rs := ?rs'; is_ps := ?pr'; is_mem := ?m' |} ?cf)
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
TF0: transl_function f = OK tf
CODE0: match_code (RTL.fn_code f) (fn_code tf)
STK0: Forall2 match_stackframe s stk'
STARSIMU0: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB0: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
H3: (fn_code tf) ! pc1 = Some x
H5: match_block (RTL.fn_code f) (fn_code tf) pc (RBnop :: RBexit None (RBgoto pc') :: nil)
Learn0: Learnt H2
x0: SeqBB.t
H0: (fn_code tf) ! pc' = Some x0
step_cf_instr tge (State stk' tf sp pc1 ?rs' ?pr' ?m') ?cf E0 ?s2
prog: RTL.program
tprog: program
ge:= Genv.globalenv prog: RTL.genv
tge:= Genv.globalenv tprog: genv
TRANSL: match_prog prog tprog
s: list RTL.stackframe
f: RTL.function
sp: val
pc: positive
rs: RTL.regset
m: mem
pc': RTL.node
H: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
stk': list stackframe
tf: function
pc1: positive
rs1: RTL.regset
m1: mem
TF: transl_function f = OK tf
CODE: forall (pc : positive) (b : SeqBB.t), (fn_code tf) ! pc = Some b -> match_block (RTL.fn_code f) (fn_code tf) pc b
STK: Forall2 match_stackframe s stk'
STARSIMU: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
x: SeqBB.t
H2: (fn_code tf) ! pc1 = Some x
H1: (RTL.fn_code f) ! pc = Some (RTL.Inop pc')
Learn: Learnt H
TF0: transl_function f = OK tf
CODE0: match_code (RTL.fn_code f) (fn_code tf)
STK0: Forall2 match_stackframe s stk'
STARSIMU0: star RTL.step ge (RTL.State s f sp pc1 rs1 m1) E0 (RTL.State s f sp pc rs m)
BB0: sem_extrap tf pc1 sp (Iexec {| is_rs := rs; is_ps := PMap.init false; is_mem := m |}) (Iexec {| is_rs := rs1; is_ps := PMap.init false; is_mem := m1 |}) (RBnop :: RBexit None (RBgoto pc') :: nil)
H3: (fn_code tf) ! pc1 = Some x
H5: match_block (RTL.fn_code f) (fn_code tf) pc (RBnop :: RBexit None (RBgoto pc') :: nil)
Learn0: Learnt H2
x0: SeqBB.t
H0: (fn_code tf) !