include "ASM/PolicyStep.ma". include alias "basics/lists/list.ma". include alias "arithmetics/nat.ma". include alias "basics/logic.ma". let rec jump_expansion_internal (program: Σl:list labelled_instruction. lt (S (|l|)) 2^16 ∧ is_well_labelled_p l) (n: ℕ) on n:(Σx:bool × (option ppc_pc_map). let 〈no_ch,pol〉 ≝ x in match pol with [ None ⇒ True | Some x ⇒ And (And (And (And (not_jump_default program x) (\fst (bvt_lookup … (bitvector_of_nat ? 0) (\snd x) 〈0,short_jump〉) = 0)) (\fst x = \fst (bvt_lookup … (bitvector_of_nat ? (|program|)) (\snd x) 〈0,short_jump〉))) (sigma_compact_unsafe program (create_label_map program) x)) (\fst x ≤ 2^16) ]) ≝ let labels ≝ create_label_map program in match n return λx.n = x → Σa:bool × (option ppc_pc_map).? with [ O ⇒ λp.〈false,pi1 ?? (jump_expansion_start program labels)〉 | S m ⇒ λp.let 〈no_ch,z〉 as p1 ≝ (pi1 ?? (jump_expansion_internal program m)) in match z return λx. z=x → Σa:bool × (option ppc_pc_map).? with [ None ⇒ λp2.〈false,None ?〉 | Some op ⇒ λp2.if no_ch then pi1 ?? (jump_expansion_internal program m) else pi1 ?? (jump_expansion_step program (pi1 ?? labels) «op,?») ] (refl … z) ] (refl … n). [5: #l #_ % | normalize nodelta cases (jump_expansion_start program (create_label_map program)) #x cases x -x [ #H % | #sigma normalize nodelta #H @conj [ @conj [ @(proj1 ?? (proj1 ?? (proj1 ?? H))) | @(proj2 ?? (proj1 ?? (proj1 ?? H))) ] | @(proj2 ?? H) ] ] | % | cases no_ch in p1; #p1 [ @(pi2 ?? (jump_expansion_internal program m)) | cases (jump_expansion_step ???) #x cases x -x #ch2 #z2 cases z2 normalize nodelta [ #_ % | #j2 #H2 @conj [ @conj [ @(proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? H2))))) | @(proj2 ?? (proj1 ?? (proj1 ?? (proj1 ?? H2)))) ] | @(proj2 ?? H2) ] ] ] | cases (jump_expansion_internal program m) in p1; #p cases p -p #p #r cases r normalize nodelta [ #_ >p2 #ABS destruct (ABS) | #map >p2 normalize nodelta #H #eq destruct (eq) @H ] ] qed. lemma pe_int_refl: ∀program.reflexive ? (sigma_jump_equal program). #program whd #x whd #n #Hn cases (bvt_lookup … (bitvector_of_nat 16 n) (\snd x) 〈0,short_jump〉) #y #z normalize nodelta @refl qed. lemma pe_int_sym: ∀program.symmetric ? (sigma_jump_equal program). #program whd #x #y #Hxy whd #n #Hn lapply (Hxy n Hn) cases (bvt_lookup … (bitvector_of_nat ? n) (\snd x) 〈0,short_jump〉) #x1 #x2 cases (bvt_lookup … (bitvector_of_nat ? n) (\snd y) 〈0,short_jump〉) #y1 #y2 normalize nodelta // qed. lemma pe_int_trans: ∀program.transitive ? (sigma_jump_equal program). #program whd #x #y #z whd in match (sigma_jump_equal ???); whd in match (sigma_jump_equal ?y ?); whd in match (sigma_jump_equal ? x z); #Hxy #Hyz #n #Hn lapply (Hxy n Hn) -Hxy lapply (Hyz n Hn) -Hyz cases (bvt_lookup … (bitvector_of_nat ? n) (\snd x) 〈0,short_jump〉) #x1 #x2 cases (bvt_lookup … (bitvector_of_nat ? n) (\snd y) 〈0,short_jump〉) #y1 #y2 cases (bvt_lookup … (bitvector_of_nat ? n) (\snd z) 〈0,short_jump〉) #z1 #z2 normalize nodelta // qed. definition policy_equal_opt ≝ λprogram:list labelled_instruction.λp1,p2:option ppc_pc_map. match p1 with [ Some x1 ⇒ match p2 with [ Some x2 ⇒ sigma_jump_equal program x1 x2 | _ ⇒ False ] | None ⇒ p2 = None ? ]. lemma pe_refl: ∀program.reflexive ? (policy_equal_opt program). #program whd #x whd cases x try % #y @pe_int_refl qed. lemma pe_sym: ∀program.symmetric ? (policy_equal_opt program). #program whd #x #y #Hxy whd cases y in Hxy; [ cases x [ #_ % | #x' #H @⊥ @(absurd ? H) /2 by nmk/ ] | #y' cases x [ #H @⊥ @(absurd ? H) whd in match (policy_equal_opt ???); @nmk #H destruct (H) | #x' #H @pe_int_sym @H ] ] qed. lemma pe_trans: ∀program.transitive ? (policy_equal_opt program). #program whd #x #y #z cases x [ #Hxy #Hyz >Hxy in Hyz; // | #x' cases y [ #H @⊥ @(absurd ? H) /2 by nmk/ | #y' cases z [ #_ #H @⊥ @(absurd ? H) /2 by nmk/ | #z' @pe_int_trans ] ] ] qed. definition step_none: ∀program.∀n. (\snd (pi1 ?? (jump_expansion_internal program n))) = None ? → ∀k.(\snd (pi1 ?? (jump_expansion_internal program (n+k)))) = None ?. #program #n lapply (refl ? (jump_expansion_internal program n)) cases (jump_expansion_internal program n) in ⊢ (???% → %); #x1 cases x1 #p1 #j1 -x1; #H1 #Heqj #Hj #k elim k [ Heqj @Hj | #k' -k Heqj2 normalize nodelta @refl | #x #H #Heqj2 #abs destruct (abs) ] ] qed. lemma jump_pc_equal: ∀program.∀n. match \snd (jump_expansion_internal program n) with [ None ⇒ True | Some p1 ⇒ match \snd (jump_expansion_internal program (S n)) with [ None ⇒ True | Some p2 ⇒ sigma_jump_equal program p1 p2 → sigma_pc_equal program p1 p2 ] ]. #program #n lapply (refl ? (jump_expansion_internal program n)) cases (jump_expansion_internal program n) in ⊢ (???% → %); #x cases x -x #Nno_ch #No cases No [ normalize nodelta #HN #NEQ @I | #Npol normalize nodelta #HN #NEQ lapply (refl ? (jump_expansion_internal program (S n))) cases (jump_expansion_internal program (S n)) in ⊢ (???% → %); #x cases x -x #Sno_ch #So cases So [ normalize nodelta #HS #SEQ @I | #Spol normalize nodelta #HS #SEQ #Hj whd in match (jump_expansion_internal program (S n)) in SEQ; (*80s*) >NEQ in SEQ; normalize nodelta cases Nno_ch in HN; [ #HN normalize nodelta #SEQ >(Some_eq ??? (proj2 ?? (pair_destruct ?????? (pi1_eq ???? SEQ)))) / by / | #HN normalize nodelta cases (jump_expansion_step ???) #x cases x -x #Stno_ch #Stno_o normalize nodelta cases Stno_o [ normalize nodelta #_ #H destruct (H) | #Stno_p normalize nodelta #HSt #STeq <(Some_eq ??? (proj2 ?? (pair_destruct ?????? (pi1_eq ???? STeq)))) in Hj; #Hj @(proj2 ?? (proj1 ?? HSt)) @(proj2 ?? (proj1 ?? (proj1 ?? HSt))) @Hj ] ] ] ] qed. lemma pe_step: ∀program:(Σl:list labelled_instruction.S (|l|) < 2^16 ∧ is_well_labelled_p l). ∀n.policy_equal_opt program (\snd (pi1 ?? (jump_expansion_internal program n))) (\snd (pi1 ?? (jump_expansion_internal program (S n)))) → policy_equal_opt program (\snd (pi1 ?? (jump_expansion_internal program (S n)))) (\snd (pi1 ?? (jump_expansion_internal program (S (S n))))). #program #n #Heq inversion (jump_expansion_internal program n) #x cases x -x #no_ch #pol cases pol normalize nodelta [ #H #Hj lapply (step_none program n) >Hj #Hn lapply (Hn (refl ??) 1) HSeq lapply (Hn (refl ??) 2) HSSeq / by / | -pol #pol #Hpol #Hn >Hn in Heq; whd in match (policy_equal_opt ???); lapply (refl ? (jump_expansion_internal program (S n))) whd in match (jump_expansion_internal program (S n)) in ⊢ (???% → ?); >Hn normalize nodelta inversion no_ch #Hno_ch normalize nodelta #Seq >Seq [ #Heq lapply (refl ? (jump_expansion_internal program (S (S n)))) whd in match (jump_expansion_internal program (S (S n))) in ⊢ (???% → ?); >Seq normalize nodelta #Teq >Teq @pe_refl | #Heq lapply (refl ? (jump_expansion_internal program (S (S n)))) whd in match (jump_expansion_internal program (S (S n))) in ⊢ (???% → ?); >Seq normalize nodelta #Teq >Teq -Teq cases (jump_expansion_step program ??) in Heq Seq ⊢ ?; (*320s*) #x cases x -x #Sno_ch #Spol normalize nodelta cases Spol [ normalize nodelta #HSn #Heq #Seq cases Heq | -Spol #Spol normalize nodelta cases Sno_ch normalize nodelta #HSn #Heq #Seq [ @pe_refl | cases daemon ] ] ] ] qed. lemma equal_remains_equal: ∀program:(Σl:list labelled_instruction. S (|l|) < 2^16 ∧ is_well_labelled_p l).∀n:ℕ. policy_equal_opt program (\snd (pi1 … (jump_expansion_internal program n))) (\snd (pi1 … (jump_expansion_internal program (S n)))) → ∀k.k ≥ n → policy_equal_opt program (\snd (pi1 … (jump_expansion_internal program n))) (\snd (pi1 … (jump_expansion_internal program k))). #program #n #Heq #k #Hk elim (le_plus_k … Hk); #z #H >H -H -Hk -k; lapply Heq -Heq; lapply n -n; elim z -z; [ #n #Heq associative_plus >(commutative_plus (S d) ?) measure_plus (commutative_plus 2 ?) @le_plus [1,3: @Hind |2,4: / by le_n/ ] ] ] qed. (* uses the second part of policy_increase *) lemma measure_incr_or_equal: ∀program:(Σl:list labelled_instruction. S (|l|) <2^16 ∧ is_well_labelled_p l). ∀policy:Σp:ppc_pc_map. (*out_of_program_none program p ∧*) not_jump_default program p ∧ \fst (bvt_lookup … (bitvector_of_nat ? 0) (\snd p) 〈0,short_jump〉) = 0 ∧ \fst p = \fst (bvt_lookup … (bitvector_of_nat ? (|program|)) (\snd p) 〈0,short_jump〉) ∧ sigma_compact_unsafe program (pi1 … (create_label_map program)) p ∧ \fst p ≤ 2^16. ∀l.|l| ≤ |program| → ∀acc:ℕ. match \snd (pi1 ?? (jump_expansion_step program (pi1 … (create_label_map program)) policy)) with [ None ⇒ True | Some p ⇒ measure_int l policy acc ≤ measure_int l p acc ]. [2: #l #_ %] #program #policy #l elim l -l; [ #Hp #acc cases (jump_expansion_step ???) #pi1 cases pi1 #p #q -pi1; cases q [ // | #x #_ @le_n ] | #h #t #Hind #Hp #acc inversion (jump_expansion_step ???) #pi1 cases pi1 -pi1 #c #r cases r [ / by I/ | #x normalize nodelta #Hx #Hjeq lapply (proj2 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hx)))) (|t|) (le_S_to_le … Hp)) whd in match (measure_int ???); whd in match (measure_int ? x ?); cases (bvt_lookup ?? (bitvector_of_nat ? (|t|)) (\snd (pi1 ?? policy)) 〈0,short_jump〉) #x1 #x2 cases (bvt_lookup ?? (bitvector_of_nat ? (|t|)) (\snd x) 〈0,short_jump〉) #y1 #y2 normalize nodelta #Hblerp cases Hblerp cases x2 cases y2 [1,4,5,7,8,9: #H cases H |2,3,6: #_ normalize nodelta [1,2: @(transitive_le ? (measure_int t x acc)) |3: @(transitive_le ? (measure_int t x (acc+1))) ] [2,4,5,6: >measure_plus [1,2: @le_plus_n_r] >measure_plus @le_plus / by le_n/] >Hjeq in Hind; #Hind @Hind @(transitive_le … Hp) @le_n_Sn |11,12,13,15,16,17: #H destruct (H) |10,14,18: normalize nodelta #_ >Hjeq in Hind; #Hind @Hind @(transitive_le … Hp) @le_n_Sn ] ] ] qed. lemma measure_full: ∀program.∀policy. measure_int program policy 0 = 2*|program| → ∀i.i<|program| → is_jump (\snd (nth i ? program 〈None ?,Comment EmptyString〉)) → (\snd (bvt_lookup ?? (bitvector_of_nat ? i) (\snd policy) 〈0,short_jump〉)) = long_jump. #program #policy elim program in ⊢ (% → ∀i.% → ? → %); [ #Hm #i #Hi @⊥ @(absurd … Hi) @not_le_Sn_O | #h #t #Hind #Hm #i #Hi #Hj cases (le_to_or_lt_eq … Hi) -Hi [ #Hi @Hind [ whd in match (measure_int ???) in Hm; cases (\snd (bvt_lookup … (bitvector_of_nat ? (|t|)) (\snd policy) 〈0,short_jump〉)) in Hm; normalize nodelta [ #H @⊥ @(absurd ? (measure_le t policy)) >H @lt_to_not_le /2 by lt_plus, le_n/ | >measure_plus >commutative_plus #H @⊥ @(absurd ? (measure_le t policy)) <(plus_to_minus … (sym_eq … H)) @lt_to_not_le normalize /2 by le_n/ | >measure_plus commutative_plus /2 by injective_plus_r/ ] | @(le_S_S_to_le … Hi) | @Hj ] | #Hi >(injective_S … Hi) whd in match (measure_int ???) in Hm; cases (\snd (bvt_lookup … (bitvector_of_nat ? (|t|)) (\snd policy) 〈0,short_jump〉)) in Hm; normalize nodelta [ #Hm @⊥ @(absurd ? (measure_le t policy)) >Hm @lt_to_not_le /2 by lt_plus, le_n/ | >measure_plus >commutative_plus #H @⊥ @(absurd ? (measure_le t policy)) <(plus_to_minus … (sym_eq … H)) @lt_to_not_le normalize /2 by le_n/ | >measure_plus commutative_plus /2 by injective_plus_r/ ] ] ] qed. (* uses second part of policy_increase *) lemma measure_special: ∀program:(Σl:list labelled_instruction. (S (|l|)) < 2^16 ∧ is_well_labelled_p l). ∀policy:Σp:ppc_pc_map. not_jump_default program p ∧ \fst (bvt_lookup … (bitvector_of_nat ? 0) (\snd p) 〈0,short_jump〉) = 0 ∧ \fst p = \fst (bvt_lookup … (bitvector_of_nat ? (|program|)) (\snd p) 〈0,short_jump〉) ∧ sigma_compact_unsafe program (pi1 … (create_label_map program)) p ∧ \fst p ≤ 2^16. match (\snd (pi1 ?? (jump_expansion_step program (pi1 … (create_label_map program)) policy))) with [ None ⇒ True | Some p ⇒ measure_int program policy 0 = measure_int program p 0 → sigma_jump_equal program policy p ]. [2: #l #_ %] #program #policy inversion (jump_expansion_step ???) #p cases p -p #ch #pol normalize nodelta cases pol [ / by I/ | #p normalize nodelta #Hpol #eqpol lapply (le_n (|program|)) @(list_ind ? (λx.|x| ≤ |pi1 ?? program| → measure_int x policy 0 = measure_int x p 0 → sigma_jump_equal x policy p) ?? (pi1 ?? program)) [ #_ #_ #i #Hi @⊥ @(absurd ? Hi) @le_to_not_lt @le_O_n | #h #t #Hind #Hp #Hm #i #Hi cases (le_to_or_lt_eq … (le_S_S_to_le … Hi)) -Hi #Hi [ @Hind [ @(transitive_le … Hp) / by / | whd in match (measure_int ???) in Hm; whd in match (measure_int ? p ?) in Hm; lapply (proj2 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hpol)))) (|t|) (le_S_to_le … Hp)) #Hinc cases (bvt_lookup ?? (bitvector_of_nat ? (|t|)) ? 〈0,short_jump〉) in Hm Hinc; #x1 #x2 cases (bvt_lookup ?? (bitvector_of_nat ? (|t|)) ? 〈0,short_jump〉); #y1 #y2 #Hm #Hinc lapply Hm -Hm; lapply Hinc -Hinc; normalize nodelta cases x2 cases y2 normalize nodelta [1: / by / |2,3: >measure_plus #_ #H @⊥ @(absurd ? (eq_plus_S_to_lt … H)) @le_to_not_lt lapply (measure_incr_or_equal program policy t ? 0) [1,3: @(transitive_le … Hp) @le_n_Sn ] >eqpol / by / |4,7,8: #H elim H #H2 [1,3,5: cases H2 |2,4,6: destruct (H2) ] |5: >measure_plus >measure_plus >commutative_plus >(commutative_plus ? 1) #_ #H @(injective_plus_r … H) |6: >measure_plus >measure_plus change with (1+1) in match (2); >assoc_plus1 >(commutative_plus 1 (measure_int ???)) #_ #H @⊥ @(absurd ? (eq_plus_S_to_lt … H)) @le_to_not_lt @monotonic_le_plus_l lapply (measure_incr_or_equal program policy t ? 0) [ @(transitive_le … Hp) @le_n_Sn ] >eqpol / by / |9: >measure_plus >measure_plus >commutative_plus >(commutative_plus ? 2) #_ #H @(injective_plus_r … H) ] | @Hi ] | >Hi whd in match (measure_int ???) in Hm; whd in match (measure_int ? p ?) in Hm; lapply (proj2 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hpol)))) (|t|) (le_S_to_le … Hp)) cases (bvt_lookup ?? (bitvector_of_nat ? (|t|)) ? 〈0,short_jump〉) in Hm; #x1 #x2 cases (bvt_lookup ?? (bitvector_of_nat ? (|t|)) ? 〈0,short_jump〉); #y1 #y2 normalize nodelta cases x2 cases y2 normalize nodelta [1,5,9: #_ #_ @refl |4,7,8: #_ #H elim H #H2 [1,3,5: cases H2 |2,4,6: destruct (H2) ] |2,3: >measure_plus #H #_ @⊥ @(absurd ? (eq_plus_S_to_lt … H)) @le_to_not_lt lapply (measure_incr_or_equal program policy t ? 0) [1,3: @(transitive_le … Hp) @le_n_Sn ] >eqpol / by / |6: >measure_plus >measure_plus change with (1+1) in match (2); >assoc_plus1 >(commutative_plus 1 (measure_int ???)) #H #_ @⊥ @(absurd ? (eq_plus_S_to_lt … H)) @le_to_not_lt @monotonic_le_plus_l lapply (measure_incr_or_equal program policy t ? 0) [ @(transitive_le … Hp) @le_n_Sn ] >eqpol / by / ] ] ] qed. lemma measure_zero: ∀l.∀program:Σl:list labelled_instruction. S (|l|) < 2^16 ∧ is_well_labelled_p l. match jump_expansion_start program (create_label_map program) with [ None ⇒ True | Some p ⇒ |l| ≤ |program| → measure_int l p 0 = 0 ]. #l #program lapply (refl ? (jump_expansion_start program (create_label_map program))) cases (jump_expansion_start program (create_label_map program)) in ⊢ (???% → %); #p #Hp #EQ cases p in Hp EQ; [ / by I/ | #pl normalize nodelta #Hpl #EQ elim l [ / by refl/ | #h #t #Hind #Hp whd in match (measure_int ???); elim (proj2 ?? (proj1 ?? Hpl) (|t|) (le_S_to_le … Hp)) #pc #Hpc >(lookup_opt_lookup_hit … Hpc 〈0,short_jump〉) normalize nodelta @Hind @(transitive_le … Hp) @le_n_Sn ] ] qed. (* the actual computation of the fixpoint *) definition je_fixpoint: ∀program:(Σl:list labelled_instruction. S (|l|) < 2^16 ∧ is_well_labelled_p l). Σp:option ppc_pc_map. match p with [ None ⇒ True | Some pol ⇒ And (And (And (\fst (bvt_lookup … (bitvector_of_nat ? 0) (\snd pol) 〈0,short_jump〉) = 0) (\fst pol = \fst (bvt_lookup … (bitvector_of_nat ? (|program|)) (\snd pol) 〈0,short_jump〉))) (sigma_compact program (pi1 … (create_label_map program)) pol)) (\fst pol ≤ 2^16) ]. #program @(\snd (jump_expansion_internal program (S (2*|program|)))) cases (dec_bounded_exists (λk.policy_equal_opt (pi1 ?? program) (\snd (pi1 ?? (jump_expansion_internal program k))) (\snd (pi1 ?? (jump_expansion_internal program (S k))))) ? (2*|program|)) [ #Hex cases Hex -Hex #k #Hk inversion (jump_expansion_internal ??) #x cases x -x #Gno_ch #Go cases Go normalize nodelta [ #H #Heq / by I/ | -Go #Gp #HGp #Geq cut (policy_equal_opt program (\snd (jump_expansion_internal program (2*|program|))) (\snd (jump_expansion_internal program (S (2*|program|))))) [ @(pe_trans … (equal_remains_equal program k (proj2 ?? Hk) (S (2*|program|)) (le_S … (le_S_to_le … (proj1 ?? Hk))))) @pe_sym @equal_remains_equal [ @(proj2 ?? Hk) | @(le_S_to_le … (proj1 ?? Hk)) ] | >Geq lapply (refl ? (jump_expansion_internal program (2*|program|))) cases (jump_expansion_internal program (2*|program|)) in ⊢ (???% → %); #x cases x -x #Fno_ch #Fo cases Fo normalize nodelta [ #H #Feq whd in match policy_equal_opt; normalize nodelta #ABS destruct (ABS) | -Fo #Fp #HFp #Feq whd in match policy_equal_opt; normalize nodelta #Heq @conj [ @conj [ @conj [ @(proj2 ?? (proj1 ?? (proj1 ?? (proj1 ?? HGp)))) | @(proj2 ?? (proj1 ?? (proj1 ?? HGp))) ] | @(equal_compact_unsafe_compact ? Fp) [ lapply (jump_pc_equal program (2*|program|)) >Feq >Geq normalize nodelta #H @H @Heq | @Heq | cases daemon (* true, but have to add this to properties *) | cases daemon | @(proj2 ?? (proj1 ?? HGp)) ] ] | @(proj2 ?? HGp) ] ] ] ] | #Hnex lapply (not_exists_forall … Hnex) -Hnex #Hfa lapply (refl ? (jump_expansion_internal program (2*|program|))) cases (jump_expansion_internal program (2*|program|)) in ⊢ (???% → %); #x cases x -x #Fno_ch #Fo cases Fo normalize nodelta [ (* None *) #HF #Feq lapply (step_none program (2*|program|) ? 1) >Feq / by / H -H normalize nodelta / by I/ | -Fo #Fp #HFp #Feq lapply (measure_full program Fp ?) [ @le_to_le_to_eq [ @measure_le | cut (∀x:ℕ.x ≤ 2*|program| → ∃p.(\snd (pi1 ?? (jump_expansion_internal program x)) = Some ? p ∧ x ≤ measure_int program p 0)) [ #x elim x [ #Hx lapply (refl ? (jump_expansion_start program (create_label_map program))) cases (jump_expansion_start program (create_label_map program)) in ⊢ (???% → %); #z cases z -z normalize nodelta [ #H #Heqn @⊥ elim (le_to_eq_plus ?? Hx) #n #Hn @(absurd … (step_none program 0 ? n)) [ whd in match (jump_expansion_internal ??); >Heqn @refl | Feq @nmk #H destruct (H) ] | #Zp #HZp #Zeq @(ex_intro ?? Zp) @conj [ whd in match (jump_expansion_internal ??); >Zeq @refl | @le_O_n ] ] | -x #x #Hind #Hx lapply (refl ? (jump_expansion_internal program (S x))) cases (jump_expansion_internal program (S x)) in ⊢ (???% → %); #z cases z -z #Sno_ch #So cases So -So [ #HSp #Seq normalize nodelta @⊥ elim (le_to_eq_plus ?? Hx) #k #Hk @(absurd … (step_none program (S x) ? k)) [ >Seq @refl | Feq @nmk #H destruct (H) ] | #Sp #HSp #Seq @(ex_intro ?? Sp) @conj [ @refl | elim (Hind (transitive_le … (le_n_Sn x) Hx)) #pol #Hpol @(le_to_lt_to_lt … (proj2 ?? Hpol)) lapply (proj1 ?? Hpol) -Hpol lapply (refl ? (jump_expansion_internal program x)) cases (jump_expansion_internal program x) in ⊢ (???% → %); #z cases z -z #Xno_ch #Xo cases Xo [ #HXp #Xeq #abs destruct (abs) | normalize nodelta #Xp #HXp #Xeq #H <(Some_eq ??? H) -H -pol lapply (Hfa x Hx) >Xeq >Seq whd in match policy_equal_opt; normalize nodelta #Hpe lapply (measure_incr_or_equal program Xp program (le_n (|program|)) 0) [ @HXp | lapply (Hfa x Hx) >Xeq >Seq lapply (measure_special program «Xp,?») [ @HXp | lapply Seq whd in match (jump_expansion_internal program (S x)); (*340s*) >Xeq normalize nodelta cases Xno_ch in HXp Xeq; #HXp #Xeq [ normalize nodelta #EQ >(proj2 ?? (pair_destruct ?????? (pi1_eq ???? EQ))) #_ #abs @⊥ @(absurd ?? abs) / by / | normalize nodelta cases (jump_expansion_step ???); #z cases z -z #stch #sto cases sto [ normalize nodelta #_ #ABS destruct (ABS) | -sto #stp normalize nodelta #Hstp #steq >(Some_eq ??? (proj2 ?? (pair_destruct ?????? (pi1_eq ???? steq)))) #Hms #Hneq #glerp elim (le_to_or_lt_eq … glerp) [ / by / | #glorp @⊥ @(absurd ?? Hneq) @Hms @glorp ] ] ] ] ] ] ] ] ] | #H elim (H (2*|program|) (le_n ?)) #plp >Feq #Hplp >(Some_eq ??? (proj1 ?? Hplp)) @(proj2 ?? Hplp) ] ] | #Hfull lapply (refl ? (jump_expansion_internal program (S (2*|program|)))) cases (jump_expansion_internal program (S (2*|program|))) in ⊢ (???% → %); #z cases z -z #Gch #Go cases Go normalize nodelta [ #HGp #Geq @I | -Go #Gp normalize nodelta #HGp #Geq @conj [ @conj [ @conj [ @(proj2 ?? (proj1 ?? (proj1 ?? (proj1 ?? HGp)))) | @(proj2 ?? (proj1 ?? (proj1 ?? HGp))) ] | @(equal_compact_unsafe_compact ? Fp) [1,2: [1: lapply (jump_pc_equal program (2*(|program|))) >Feq >Geq normalize nodelta #H @H ] #i #Hi inversion (is_jump (\snd (nth i ? program 〈None ?, Comment EmptyString〉))) [1,3: #Hj whd in match (jump_expansion_internal program (S (2*|program|))) in Geq; (*85s*) >Feq in Geq; normalize nodelta cases Fno_ch [1,3: normalize nodelta #Heq >(Some_eq ??? (proj2 ?? (pair_destruct ?????? (pi1_eq ???? Heq)))) % |2,4: normalize nodelta cases (jump_expansion_step ???) #x cases x -x #stch #sto normalize nodelta cases sto [1,3: normalize nodelta #_ #X destruct (X) |2,4: -sto #stp normalize nodelta #Hst #Heq <(Some_eq ??? (proj2 ?? (pair_destruct ?????? (pi1_eq ???? Heq)))) lapply (proj2 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hst)))) i (le_S_to_le … Hi)) lapply (Hfull i Hi ?) [1,3: >Hj %] cases (bvt_lookup … (bitvector_of_nat ? i) (\snd Fp) 〈0,short_jump〉) #fp #fj #Hfj >Hfj normalize nodelta cases (bvt_lookup … (bitvector_of_nat ? i) (\snd stp) 〈0,short_jump〉) #stp #stj cases stj normalize nodelta [1,2,4,5: #H cases H #H2 cases H2 destruct (H2) |3,6: #_ @refl ] ] ] |2,4: #Hj >(proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? HGp))) i Hi ?) [2,4:>Hj %] >(proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? HFp))) i Hi ?) [2,4:>Hj] % ] | cases daemon (* true, but have to add to properties in some way *) | cases daemon | @(proj2 ?? (proj1 ?? HGp)) ] ] | @(proj2 ?? HGp) ] ] ] ] | #n cases (jump_expansion_internal program n) cases (jump_expansion_internal program (S n)) #x cases x -x #nch #npol normalize nodelta #Hnpol #x cases x -x #Sch #Spol normalize nodelta #HSpol cases npol in Hnpol; cases Spol in HSpol; [ #Hnpol #HSpol %1 // |2,3: #x #Hnpol #HSpol %2 @nmk whd in match (policy_equal_opt ???); // #H destruct (H) |4: #np #Hnp #Sp #HSp whd in match (policy_equal_opt ???); @dec_bounded_forall #m cases (bvt_lookup ?? (bitvector_of_nat 16 m) ? 〈0,short_jump〉) #x1 #x2 cases (bvt_lookup ?? (bitvector_of_nat ? m) ? 〈0,short_jump〉) #y1 #y2 normalize nodelta @dec_eq_jump_length ] ] qed. include alias "arithmetics/nat.ma". include alias "basics/logic.ma". lemma pc_increases: ∀i,j:ℕ.∀program.∀pol:Σp:ppc_pc_map. And (And (And (\fst (bvt_lookup … (bitvector_of_nat ? 0) (\snd p) 〈0,short_jump〉) = 0) (\fst p = \fst (bvt_lookup … (bitvector_of_nat ? (|program|)) (\snd p) 〈0,short_jump〉))) (sigma_compact program (create_label_map program) p)) (\fst p ≤ 2^16).i ≤ j → j ≤ |program| → \fst (bvt_lookup (ℕ×jump_length) 16 (bitvector_of_nat 16 i) (\snd pol) 〈0,short_jump〉) ≤ \fst (bvt_lookup (ℕ×jump_length) 16 (bitvector_of_nat 16 j) (\snd pol) 〈0,short_jump〉). #i #j #program #pol #H elim (le_to_eq_plus … H) #n #Hn >Hn -Hn -j elim n [ (lookup_opt_lookup_hit … EQ 〈0,short_jump〉) >(lookup_opt_lookup_hit … SEQ 〈0,short_jump〉) >Hcomp @le_plus_n_r ] ] ] qed. (* The glue between Policy and Assembly. *) definition jump_expansion': ∀program:preamble × (Σl:list labelled_instruction.S (|l|) < 2^16 ∧ is_well_labelled_p l). option (Σsigma_policy:(Word → Word) × (Word → bool). let 〈sigma,policy〉≝ sigma_policy in sigma_policy_specification 〈\fst program,\snd program〉 sigma policy) ≝ λprogram. let f: option ppc_pc_map ≝ je_fixpoint (\snd program) in match f return λx.f = x → ? with [ None ⇒ λp.None ? | Some x ⇒ λp.Some ? «〈(λppc.let pc ≝ \fst (bvt_lookup ?? ppc (\snd x) 〈0,short_jump〉) in bitvector_of_nat 16 pc), (λppc.let jl ≝ \snd (bvt_lookup ?? ppc (\snd x) 〈0,short_jump〉) in jmpeqb jl long_jump)〉,?» ] (refl ? f). normalize nodelta in p; whd in match sigma_policy_specification; normalize nodelta lapply (pi2 ?? (je_fixpoint (\snd program))) >p normalize nodelta cases x #fpc #fpol #Hfpol cases Hfpol ** #Hfpol1 #Hfpol2 #Hfpol3 #Hfpol4 @conj [ >Hfpol1 % | #ppc #ppc_ok normalize nodelta >(?:\fst (fetch_pseudo_instruction (pi1 … (\snd program)) ppc ppc_ok) = \snd (nth (nat_of_bitvector … ppc) ? (\snd program) 〈None ?, Comment EmptyString〉)) [2: whd in match fetch_pseudo_instruction; normalize nodelta >(nth_safe_nth … 〈None ?, Comment EmptyString〉) cases (nth (nat_of_bitvector ? ppc) ? (\snd program) 〈None ?, Comment EmptyString〉) #lbl #ins % ] lapply (Hfpol3 ? (nat_of_bitvector ? ppc) ppc_ok) [2: >bitvector_of_nat_inverse_nat_of_bitvector inversion (lookup_opt ????) normalize nodelta [ #Hl #abs cases abs ] * #pc #jl #Hl normalize nodelta inversion (lookup_opt ????) normalize nodelta [ #Hl #abs cases abs ] * #Spc #Sjl #SHL lapply SHL bitvector_of_nat_inverse_nat_of_bitvector >add_commutative #SHl normalize nodelta #Hcompact @conj [ >(lookup_opt_lookup_hit … SHl 〈0,short_jump〉) >(lookup_opt_lookup_hit … Hl 〈0,short_jump〉) >add_bitvector_of_nat_plus >Hcompact % | (* Basic proof scheme: - ppc < |snd program|, hence our instruction is in the program - either we are the last non-zero-size instruction, in which case we are either smaller than 2^16 (because the entire program is), or we are exactly 2^16 and something weird happens - or we are not, in which case we are definitely smaller than 2^16 (by transitivity through the next non-zero instruction) *) elim (le_to_or_lt_eq … Hfpol4) #Hfpc [ %1 @(le_to_lt_to_lt … Hfpc) >Hfpol2 >(lookup_opt_lookup_hit … Hl 〈0,short_jump〉) >nat_of_bitvector_bitvector_of_nat_inverse [2: lapply (pc_increases (nat_of_bitvector 16 ppc) (|\snd program|) (\snd program) «〈fpc,fpol〉,Hfpol» (le_S_to_le … ppc_ok) (le_n ?)) >bitvector_of_nat_inverse_nat_of_bitvector >(lookup_opt_lookup_hit … Hl 〈0,short_jump〉) #H @(le_to_lt_to_lt … Hfpc) >Hfpol2 @H ] lapply (pc_increases (S (nat_of_bitvector 16 ppc)) (|\snd program|) (\snd program) «〈fpc,fpol〉,Hfpol» ppc_ok (le_n ?)) >(lookup_opt_lookup_hit … SHL 〈0,short_jump〉) >Hcompact #X @X | (* the program is of length 2^16 and ppc is followed by only zero-size instructions * until the end of the program *) elim (le_to_or_lt_eq … (pc_increases (nat_of_bitvector ? ppc) (|\snd program|) (\snd program) «〈fpc,fpol〉,Hfpol» (le_S_to_le … ppc_ok) (le_n ?))) [ >bitvector_of_nat_inverse_nat_of_bitvector >(lookup_opt_lookup_hit … Hl 〈0,short_jump〉) #Hpc normalize nodelta >nat_of_bitvector_bitvector_of_nat_inverse [2: Hfpol2 @Hpc ] elim (le_to_or_lt_eq … (pc_increases (S (nat_of_bitvector ? ppc)) (|\snd program|) (\snd program) «〈fpc,fpol〉,Hfpol» ppc_ok (le_n ?))) Hfpc >(lookup_opt_lookup_hit … SHL 〈0,short_jump〉) #HSpc [ %1 >Hcompact in HSpc; #X @X | %2 @conj [2: >Hcompact in HSpc; #X @X | #ppc' #ppc_ok' #Hppc' (* S ppc < ppc' < |\snd program| *) (* lookup S ppc = 2^16 *) (* lookup |\snd program| = 2^16 *) (* lookup ppc' = 2^16 → instruction size = 0 *) lapply (Hfpol3 ? (nat_of_bitvector ? ppc') ppc_ok') [2: >bitvector_of_nat_inverse_nat_of_bitvector inversion (lookup_opt ????) normalize nodelta [ #_ #abs cases abs | * #xpc #xjl #XEQ normalize nodelta inversion (lookup_opt ????) normalize nodelta [ #_ #abs cases abs | * #Sxpc #Sxjl #SXEQ normalize nodelta #Hpcompact lapply (pc_increases (S (nat_of_bitvector ? ppc)) (nat_of_bitvector ? ppc') (\snd program) «〈fpc,fpol〉,Hfpol» Hppc' (le_S_to_le … ppc_ok')) >(lookup_opt_lookup_hit … SHL 〈0,short_jump〉) >HSpc #Hle1 lapply (pc_increases (nat_of_bitvector ? ppc') (|\snd program|) (\snd program) «〈fpc,fpol〉,Hfpol» (le_S_to_le … ppc_ok') (le_n ?)) Hfpc #Hle2 lapply (le_to_le_to_eq ?? Hle2 Hle1) -Hle2 -Hle1 >bitvector_of_nat_inverse_nat_of_bitvector >(lookup_opt_lookup_hit … XEQ 〈0,short_jump〉) #Hxpc lapply (pc_increases (S (nat_of_bitvector ? ppc)) (S (nat_of_bitvector ? ppc')) (\snd program) «〈fpc,fpol〉,Hfpol» (le_S_to_le … (le_S_S … Hppc')) ppc_ok') >(lookup_opt_lookup_hit … SHL 〈0,short_jump〉) >HSpc #Hle1 lapply (pc_increases (S (nat_of_bitvector ? ppc')) (|\snd program|) (\snd program) «〈fpc,fpol〉,Hfpol» ppc_ok' (le_n ?)) Hfpc #Hle2 lapply (le_to_le_to_eq ?? Hle2 Hle1) -Hle1 -Hle2 >(lookup_opt_lookup_hit … SXEQ 〈0,short_jump〉) #HSxpc >Hxpc in Hpcompact; >HSxpc whd in match create_label_map; #H @(plus_equals_zero (2^16)) whd in match fetch_pseudo_instruction; normalize nodelta >(nth_safe_nth … 〈None ?, Comment EmptyString〉) cases (nth (nat_of_bitvector ? ppc') ? (\snd program) 〈None ?, Comment EmptyString〉) in H; #lbl #ins normalize nodelta #X @sym_eq @X ] ] ] ] ] | >bitvector_of_nat_inverse_nat_of_bitvector Hfpc >(lookup_opt_lookup_hit … Hl 〈0,short_jump〉) #Hpc %1 >Hpc >bitvector_of_nat_exp_zero whd in match (nat_of_bitvector ? (zero ?)); Hli lapply (assembly1_pseudoinstruction_lt_2_to_16 ??? ppc ??) [6: >Hass / by / ] ] ] ] ] ] qed.