[sledge sem] Rework emap invariant

Summary:
The LLVM->LLAIR translation keeps a mapping of variables to expressions.
Previously, the invariants related to that mapping were kept in the
state relation, and so the proof needed show that they were preserved
along execution traces. This wasn't obvious as the state changes in
non-SSA ways during evaluation, but the correctness of the mappings is
heavily based on the program being in SSA form. This change separates
out the invariants, and the proof uses the final mapping that the
compiler builds, which contains all of the relevant bindings that might
be needed during execution.

Reviewed By: jberdine

Differential Revision: D20625109

fbshipit-source-id: d4c2dfe19
master
Scott Owens 5 years ago committed by Facebook GitHub Bot
parent 30ca51366d
commit 2a6ababa99

@ -828,9 +828,17 @@ Definition prog_ok_def:
((* All non-entry blocks have a proper header, and entry blocks don't *) ((* All non-entry blocks have a proper header, and entry blocks don't *)
∀fname dec. ∀fname dec.
alookup p fname = Some dec alookup p fname = Some dec
(every (\b. fst b = None (snd b).h = Entry) dec.blocks)) every (\b. fst b = None (snd b).h = Entry) dec.blocks)
((* The blocks in a definition have distinct labels.*)
∀fname dec.
alookup p fname = Some dec
all_distinct (map fst dec.blocks))
(* There is a main function *) (* There is a main function *)
∃dec. alookup p (Fn "main") = Some dec (∃dec. alookup p (Fn "main") = Some dec)
(* No phi instruction assigns the same register twice *)
(∀ip from_l phis.
get_instr p ip (Inr (from_l, phis))
all_distinct (map (λp. case p of Phi r _ _ => r) phis))
End End
(* All call frames have a good return address, and the stack allocations of the (* All call frames have a good return address, and the stack allocations of the

@ -195,26 +195,40 @@ Inductive uses:
uses prog ip r) uses prog ip r)
End End
Definition cidx_to_num_def:
(cidx_to_num (IntC _ n) = Num (ABS n))
(cidx_to_num _ = 0)
End
(* Convert index lists as for GEP into number lists, for the purpose of
* calculating types. Everything goes to 0 but for positive integer constants,
* because those things can't be used to index anything but arrays, and the type
* for the array contents doesn't depend on the index's value. *)
Definition idx_to_num_def:
(idx_to_num (_, (Constant (IntC _ n))) = Num (ABS n))
(idx_to_num (_, _) = 0)
End
(* The registers that an instruction assigns *) (* The registers that an instruction assigns *)
Definition instr_assigns_def: Definition instr_assigns_def:
(instr_assigns (Invoke r _ _ _ _ _) = {r}) (instr_assigns (Invoke r t _ _ _ _) = {(r,t)})
(instr_assigns (Sub r _ _ _ _ _) = {r}) (instr_assigns (Sub r _ _ t _ _) = {(r,t)})
(instr_assigns (Extractvalue r _ _) = {r}) (instr_assigns (Extractvalue r (t,_) idx) = {(r,THE (extract_type t (map cidx_to_num idx)))})
(instr_assigns (Insertvalue r _ _ _) = {r}) (instr_assigns (Insertvalue r (t,_) _ _) = {(r, t)})
(instr_assigns (Alloca r _ _) = {r}) (instr_assigns (Alloca r t _) = {(r,PtrT t)})
(instr_assigns (Load r _ _) = {r}) (instr_assigns (Load r t _) = {(r,t)})
(instr_assigns (Gep r _ _ _) = {r}) (instr_assigns (Gep r t _ idx) = {(r,PtrT (THE (extract_type t (map idx_to_num idx))))})
(instr_assigns (Cast r _ _ _) = {r}) (instr_assigns (Cast r _ _ t) = {(r,t)})
(instr_assigns (Icmp r _ _ _ _) = {r}) (instr_assigns (Icmp r _ _ _ _) = {(r, IntT W1)})
(instr_assigns (Call r _ _ _) = {r}) (instr_assigns (Call r t _ _) = {(r,t)})
(instr_assigns (Cxa_allocate_exn r _) = {r}) (instr_assigns (Cxa_allocate_exn r _) = {(r,ARB)})
(instr_assigns (Cxa_begin_catch r _) = {r}) (instr_assigns (Cxa_begin_catch r _) = {(r,ARB)})
(instr_assigns (Cxa_get_exception_ptr r _) = {r}) (instr_assigns (Cxa_get_exception_ptr r _) = {(r,ARB)})
(instr_assigns _ = {}) (instr_assigns _ = {})
End End
Definition phi_assigns_def: Definition phi_assigns_def:
phi_assigns (Phi r _ _) = r phi_assigns (Phi r t _) = (r,t)
End End
Inductive assigns: Inductive assigns:
@ -256,11 +270,13 @@ Definition is_ssa_def:
(∀fname. (∀fname.
(* No register is assigned in two different instructions *) (* No register is assigned in two different instructions *)
(∀r ip1 ip2. (∀r ip1 ip2.
r assigns prog ip1 r assigns prog ip2 ip1.f = fname ip2.f = fname r image fst (assigns prog ip1) r image fst (assigns prog ip2)
ip1.f = fname ip2.f = fname
ip1 = ip2)) ip1 = ip2))
(* Each use is dominated by its assignment *) (* Each use is dominated by its assignment *)
(∀ip1 r. r uses prog ip1 ∃ip2. ip2.f = ip1.f r assigns prog ip2 dominates prog ip2 ip1) (∀ip1 r. r uses prog ip1
∃ip2. ip2.f = ip1.f r image fst (assigns prog ip2) dominates prog ip2 ip1)
End End
Theorem dominates_trans: Theorem dominates_trans:
@ -389,7 +405,7 @@ Definition live_def:
{ r | ∃path. { r | ∃path.
good_path prog (ip::path) good_path prog (ip::path)
r uses prog (last (ip::path)) r uses prog (last (ip::path))
∀ip2. ip2 set (front (ip::path)) r assigns prog ip2 } ∀ip2. ip2 set (front (ip::path)) r image fst (assigns prog ip2) }
End End
Theorem get_instr_live: Theorem get_instr_live:
@ -412,7 +428,7 @@ QED
Theorem live_gen_kill: Theorem live_gen_kill:
∀prog ip ip'. ∀prog ip ip'.
live prog ip = live prog ip =
BIGUNION {live prog ip' | ip' | ip' next_ips prog ip} DIFF assigns prog ip uses prog ip BIGUNION {live prog ip' | ip' | ip' next_ips prog ip} DIFF image fst (assigns prog ip) uses prog ip
Proof Proof
rw [live_def, EXTENSION] >> eq_tac >> rw [] rw [live_def, EXTENSION] >> eq_tac >> rw []
>- ( >- (
@ -433,7 +449,7 @@ QED
Theorem ssa_dominates_live_range_lem: Theorem ssa_dominates_live_range_lem:
∀prog r ip1 ip2. ∀prog r ip1 ip2.
is_ssa prog ip1.f = ip2.f r assigns prog ip1 r live prog ip2 is_ssa prog ip1.f = ip2.f r image fst (assigns prog ip1) r live prog ip2
dominates prog ip1 ip2 dominates prog ip1 ip2
Proof Proof
rw [dominates_def, is_ssa_def, live_def] >> rw [dominates_def, is_ssa_def, live_def] >>
@ -492,7 +508,7 @@ Theorem ssa_dominates_live_range:
∀prog r ip. ∀prog r ip.
is_ssa prog r uses prog ip is_ssa prog r uses prog ip
∃ip1. ip1.f = ip.f r assigns prog ip1 ∃ip1. ip1.f = ip.f r image fst (assigns prog ip1)
∀ip2. ip2.f = ip.f r live prog ip2 ∀ip2. ip2.f = ip.f r live prog ip2
dominates prog ip1 ip2 dominates prog ip1 ip2
Proof Proof

@ -9,7 +9,7 @@
open HolKernel boolLib bossLib Parse; open HolKernel boolLib bossLib Parse;
open arithmeticTheory pred_setTheory; open arithmeticTheory pred_setTheory;
open settingsTheory llvmTheory llairTheory; open settingsTheory llvmTheory llvm_ssaTheory llairTheory;
new_theory "llvm_to_llair"; new_theory "llvm_to_llair";
@ -230,25 +230,11 @@ Datatype:
| Call | Call
End End
(* Convert index lists as for GEP into number lists, for the purpose of
* calculating types. Everything goes to 0 but for positive integer constants,
* because those things can't be used to index anything but arrays, and the type
* for the array contents doesn't depend on the index's value. *)
Definition idx_to_num_def:
(idx_to_num (_, (Constant (IntC _ n))) = Num (ABS n))
(idx_to_num (_, _) = 0)
End
Definition cidx_to_num_def:
(cidx_to_num (IntC _ n) = Num (ABS n))
(cidx_to_num _ = 0)
End
Definition classify_instr_def: Definition classify_instr_def:
(classify_instr (Call _ _ _ _) = Call) (classify_instr (Call _ _ _ _) = Call)
(classify_instr (Ret _) = Term) (classify_instr (Ret _) = Term)
(classify_instr (Br _ _ _) = Term) (classify_instr (Br _ _ _) = Term)
(classify_instr (Invoke _ _ _ _ _ _) = Term) (classify_instr (Invoke _ _ _ _ _ _) = Call)
(classify_instr Unreachable = Term) (classify_instr Unreachable = Term)
(classify_instr (Exit _) = Term) (classify_instr (Exit _) = Term)
(classify_instr (Load _ _ _) = Non_exp) (classify_instr (Load _ _ _) = Non_exp)
@ -273,10 +259,11 @@ End
Definition extend_emap_non_exp_def: Definition extend_emap_non_exp_def:
(extend_emap_non_exp emap (Load r t _) = emap |+ (r, Var (translate_reg r t) F)) (extend_emap_non_exp emap (Load r t _) = emap |+ (r, Var (translate_reg r t) F))
(extend_emap_non_exp emap (Call r t _ _) = emap |+ (r, Var (translate_reg r t) F)) (extend_emap_non_exp emap (Call r t _ _) = emap |+ (r, Var (translate_reg r t) F))
(extend_emap_non_exp emap (Invoke r t _ _ _ _) = emap |+ (r, Var (translate_reg r t) F))
(extend_emap_non_exp emap (Cxa_begin_catch r _) = emap |+ (r, Var (translate_reg r ARB) F))
(extend_emap_non_exp emap _ = emap) (extend_emap_non_exp emap _ = emap)
End End
(* Given a non-empty list of blocks, add an inst to the first one *) (* Given a non-empty list of blocks, add an inst to the first one *)
Definition add_to_first_block_def: Definition add_to_first_block_def:
(add_to_first_block i [] = []) (add_to_first_block i [] = [])

File diff suppressed because it is too large Load Diff

@ -942,6 +942,12 @@ Proof
metis_tac [] metis_tac []
QED QED
Theorem union_diff:
!s1 s2 s3. s1 s2 DIFF s3 = (s1 DIFF s3) (s2 DIFF s3)
Proof
rw [EXTENSION] >> metis_tac []
QED
(* ----- finite map theorems ----- *) (* ----- finite map theorems ----- *)
Theorem drestrict_fupdate_list[simp]: Theorem drestrict_fupdate_list[simp]:
@ -950,4 +956,13 @@ Proof
Induct_on `l` >> rw [FUPDATE_LIST_THM] >> pairarg_tac >> fs [] Induct_on `l` >> rw [FUPDATE_LIST_THM] >> pairarg_tac >> fs []
QED QED
Theorem fupdate_list_elim:
∀m l. (∀k v. mem (k,v) l flookup m k = Some v) m |++ l = m
Proof
Induct_on `l` >> rw [FUPDATE_LIST_THM] >>
rename1 `_ |+ kv |++ _ = _` >>
`m |+ kv = m` by (PairCases_on `kv` >> irule FUPDATE_ELIM >> fs [FLOOKUP_DEF]) >>
rw []
QED
export_theory (); export_theory ();

Loading…
Cancel
Save