include "joint/TranslateUtils.ma". include "common/extraGlobalenvs.ma". include "utilities/hide.ma". definition graph_to_lin_statement : ∀p : unserialized_params.∀globals. joint_statement (mk_graph_params p) globals → joint_statement (mk_lin_params p) globals ≝ λp,globals,stmt.match stmt return λ_.joint_statement (mk_lin_params ?) ? with [ sequential c _ ⇒ sequential … c it | final c ⇒ final … c ]. lemma graph_to_lin_labels : ∀p : unserialized_params.∀globals,s. stmt_labels … (graph_to_lin_statement p globals s) = stmt_explicit_labels … s. #p#globals * [//] * // qed. (* discard all elements passing test, return first element failing it *) (* and the rest of the list, if any. *) let rec chop A (test : A → bool) (l : list A) on l : option (A × (list A)) ≝ match l with [ nil ⇒ None ? | cons x l' ⇒ if test x then chop A test l' else return 〈x, l'〉 ]. lemma chop_ok : ∀A,f,l. match chop A f l with [ Some pr ⇒ let x ≝ \fst pr in let post ≝ \snd pr in ∃pre.All ? (λx.bool_to_Prop (f x)) pre ∧ l = pre @ x :: post ∧ ¬bool_to_Prop (f x) | None ⇒ All A (λx.bool_to_Prop (f x)) l ]. #A #f #l elim l [ % | #hd #tl #IH whd in match (chop ???); elim (true_or_false_Prop (f hd)) #H >H normalize nodelta [ lapply IH elim (chop ???) normalize nodelta [ #G %{H G} | #pr * #pre ** #H1 #H2 #H3 %{(hd :: pre)} %{H3} % [ %{H H1} | >H2 % ] ] | %{[ ]} %{H} %% ] ] qed. unification hint 0 ≔ p, globals; lp ≟ lin_params_to_params p, sp ≟ stmt_pars lp, js ≟ joint_statement sp globals, lo ≟ labelled_obj LabelTag js (*----------------------------*)⊢ list lo ≡ codeT lp globals. definition graph_visit_ret_type ≝ λp,globals.λg : codeT (mk_graph_params p) globals. λentry : label. (Σ〈visited' : identifier_map LabelTag ℕ, required' : identifier_set LabelTag, generated' : codeT (mk_lin_params p) globals〉.'hide ( And (And (And (And (lookup ?? visited' entry = Some ? 0) (required' ⊆ visited')) (∃last.stmt_at … generated' 0 = Some ? (final … last))) (code_forall_labels … (λl.bool_to_Prop (l∈required')) (rev … generated'))) (∀l,n.lookup ?? visited' l = Some ? n → And (bool_to_Prop (code_has_label … (rev ? generated') l)) (∃s.And (And (lookup … g l = Some ? s) (nth_opt … n (rev … generated') = Some ? 〈Some ? l, graph_to_lin_statement … s〉)) (opt_All ? (λnext.Or (lookup … visited' next = Some ? (S n)) (nth_opt … (S n) (rev … generated') = Some ? 〈None ?, GOTO … next〉)) (stmt_implicit_label … s)))))). unification hint 0 ≔ tag ⊢ identifier_set tag ≡ identifier_map tag unit. include alias "common/Identifiers.ma". lemma lookup_safe_elim : ∀tag,A.∀P : A → Prop.∀m,i,prf. (∀x.lookup tag A m i = Some ? x → P x) → P (lookup_safe tag A m i prf). #tag #A #P #m #i #prf #H @H @lookup_eq_safe qed. let rec graph_visit (p : unserialized_params) (globals: list ident) (g : codeT (mk_graph_params p) globals) (* = graph (joint_statement (mk_graph_params p) globals *) (required: identifier_set LabelTag) (visited: identifier_map LabelTag ℕ) (* the reversed index of labels in the new code *) (generated: codeT (mk_lin_params p) globals) (* ≝ list ((option label)×(joint_statement (mk_lin_params p) globals)) *) (visiting: list label) (gen_length : ℕ) (n: nat) (entry : label) (g_prf : code_closed … g) (required_prf1 : ∀i.i∈required → Or (In ? visiting i) (bool_to_Prop (i∈visited))) (required_prf2 : code_forall_labels … (λl.bool_to_Prop (l ∈ required)) (rev … generated)) (generated_prf1 : ∀l,n.lookup … visited l = Some ? n → hide_Prop ( And (bool_to_Prop (code_has_label … (rev ? generated) l)) (∃s.And (And (lookup … g l = Some ? s) (nth_opt ? n (rev … generated) = Some ? 〈Some ? l, graph_to_lin_statement … s〉)) (opt_All ? (λnext.match lookup … visited next with [ Some n' ⇒ Or (n' = S n) (nth_opt ? (S n) (rev ? generated) = Some ? 〈None ?, GOTO … next〉) | None ⇒ And (nth_opt ? 0 visiting = Some ? next) (S n = gen_length) ]) (stmt_implicit_label … s))))) (generated_prf2 : ∀l.lookup … visited l = None ? → does_not_occur … l (rev ? generated)) (generated_prf3 : (∃last.stmt_at … generated 0 = Some ? (final … last)) ∨ (¬All ? (λx.bool_to_Prop (x∈visited)) visiting)) (visiting_prf : All … (λl.bool_to_Prop (l∈g)) visiting) (gen_length_prf : gen_length = length ? generated) (entry_prf : Or (And (And (visiting = [entry]) (gen_length = 0)) (Not (bool_to_Prop (entry∈visited)))) (lookup … visited entry = Some ? 0)) (n_prf : le (id_map_size … g) (plus n (id_map_size … visited))) on n : graph_visit_ret_type … g entry ≝ match chop ? (λx.x∈visited) visiting return λx.? → graph_visit_ret_type … g entry with [ None ⇒ λH. «〈visited, required, generated〉, ?» | Some pr ⇒ λH. let vis_hd ≝ \fst pr in let vis_tl ≝ \snd pr in match n return λx.? → graph_visit_ret_type … g entry with [ O ⇒ λn_prf'.⊥ | S n' ⇒ λn_prf'. (* add the label to the visited ones *) let visited' ≝ add … visited vis_hd gen_length in (* take l's statement *) let hd_vis_in_g ≝ (hide_prf ? ?) in let statement ≝ lookup_safe ?? g vis_hd hd_vis_in_g in (* translate it to its linear version *) let translated_statement ≝ graph_to_lin_statement p globals statement in (* add the translated statement to the code (which will be reversed later) *) let generated' ≝ 〈Some … vis_hd, translated_statement〉 :: generated in let required' ≝ union_set ??? (set_from_list … (stmt_explicit_labels … statement)) required in (* put successors in the visiting worklist *) let visiting' ≝ stmt_labels … statement @ vis_tl in (* finally, check the implicit successor *) (* if it has been already visited, we need to add a GOTO *) let add_req_gen ≝ match stmt_implicit_label … statement with [ Some next ⇒ if next ∈ visited' then 〈1, {(next)}, [〈None label, (GOTO … next : joint_statement ??)〉]〉 else 〈0, ∅, [ ]〉 | None ⇒ 〈0, ∅, [ ]〉 ] in (* prepare a common utility to deal with add_req_gen *) let add_req_gen_prf : ∀P : (ℕ × (identifier_set LabelTag) × (codeT (mk_lin_params p) globals)) → Prop. (opt_All ? (λnext.¬bool_to_Prop (next ∈ visited')) (stmt_implicit_label … statement) → P 〈0,∅,[ ]〉) → (∀next.stmt_implicit_label … statement = Some ? next → next ∈ visited' → P 〈1, {(next)}, [〈None ?, GOTO (mk_lin_params p) next〉]〉) → P add_req_gen ≝ hide_prf ?? in graph_visit ??? (union_set ??? (\snd (\fst add_req_gen)) required') visited' (\snd add_req_gen @ generated') visiting' (plus (\fst (\fst add_req_gen)) (S gen_length)) n' entry g_prf ????????? ] n_prf ] (chop_ok ? (λx.x∈visited) visiting). (*cases daemon qed. *) whd [ (* base case, visiting is all visited *) %[%[%[%]]] [ elim entry_prf [ ** #eq_visiting #gen_length_O #entry_vis >eq_visiting in H; * >entry_vis * | // ] | #l #l_req elim (required_prf1 … l_req) #G [ @(All_In … H G) | assumption ] | cases generated_prf3 [//] * #ABS @⊥ @ABS assumption | assumption | #l #n #H elim (generated_prf1 … H) #H1 * #s ** #H2 #H3 #H4 % [assumption] %{s} % [% assumption | @(opt_All_mp … H4) -l #l lapply (in_map_domain … visited l) elim (true_or_false_Prop (l∈visited)) #l_vis >l_vis normalize nodelta [ * #n' ] #EQlookup >EQlookup normalize nodelta * [ #EQn' % >EQn' % | #H %2{H} | #H' lapply (All_nth … H … H') >l_vis * ] ] ] (* first, close goals where add_gen_req plays no role *) |13: (* vis_hd is in g *) elim H #pre ** #_ #H2 #_ @(All_split … visiting_prf … H2) |2: (* n = 0 absrud *) elim H #pre ** #_ #H2 #H3 @(absurd … n_prf') @lt_to_not_le lapply (add_size … visited vis_hd 0 (* dummy value *)) >H3 normalize nodelta whd in match (lt ? ?); whd in match (1 + ?); #EQ (lookup_eq_safe … H) % | #l #H elim (generated_prf1 … (lookup_eq_safe … H)) #_ * #s ** #s_in_g #_ #_ whd in ⊢ (?%); >s_in_g % ] |8: elim H #pre ** #_ #H2 #_ @All_append [ elim(g_prf … vis_hd statement ?) [2:@lookup_eq_safe] #G1 #G2 @(All_append … G1) whd in G2; lapply G2 elim statement normalize nodelta #s [2: #_ %] #l #G' %{G'} % | >H2 in visiting_prf; #H' lapply (All_append_r … H') -H' * #_ // ] |10: elim H #pre ** #_ #H2 #H3 -add_req_gen_prf %2 elim entry_prf [ ** >H2 cases pre [2: #x' #pre' #ABS normalize in ABS; destruct(ABS) cases pre' in e0; [2: #x'' #pre''] #ABS' normalize in ABS'; destruct(ABS') |1: #EQ normalize in EQ; destruct(EQ) #eq_gen_length #_ >lookup_add_hit >eq_gen_length % ] | #lookup_entry cut (entry ≠ vis_hd) [ % whd in match vis_hd; #entry_vis_hd lookup_entry * #ABS @ABS % ] #entry_not_vis_hd >(lookup_add_miss ?????? entry_not_vis_hd) assumption ] |11: elim H #pre ** #_ #_ #H3 >commutative_plus >add_size >H3 normalize nodelta whd in match (1 + ?); >commutative_plus assumption |12: (* add_req_gen utility *) #P whd in match add_req_gen; elim (stmt_implicit_label ???) [ #H #_ @H % ] #next normalize nodelta elim (true_or_false_Prop (next ∈ visited')) #next_vis >next_vis normalize nodelta [ #_ #H @H [% | assumption] | #H #_ @H whd >next_vis % * ] | elim H #pre ** #H1 #H2 #_ #i >mem_set_union #H elim (orb_Prop_true … H) -H [ @add_req_gen_prf [ #_ >mem_set_empty * ] #next #_ #next_vis #H >(mem_set_singl_to_eq … H) %2 assumption | >mem_set_union #H elim (orb_Prop_true … H) -H [ #i_expl %1 @Exists_append_l @Exists_append_r @mem_list_as_set @i_expl | (* i was already required *) #i_req elim (required_prf1 … i_req) [ >H2 #H elim (Exists_append … H) -H [ (* i in preamble → i∈visited *) #i_pre %2 >mem_set_add @orb_Prop_r lapply (All_In … H1 i_pre) #H @H | * [ (* i is vis_hd *) #eq_i >eq_i %2 @mem_set_add_id | (* i in vis_tl → i∈visiting' *) #i_post % @Exists_append_r assumption ] ] | (* i in visited *) #i_vis %2 >mem_set_add @orb_Prop_r assumption ] ] ] |4,5,6: change with reverse in match rev; >reverse_append whd in match (reverse ??); >rev_append_def >associative_append [ #pt #s @(leb_elim (S pt) (|generated|)) #cmp whd in match (stmt_at ????); [ >nth_opt_append_l [2: >length_reverse assumption ] change with (stmt_at ???? = ? → ?) #EQ lapply(required_prf2 … EQ) @All_mp #l #l_req >mem_set_union @orb_Prop_r >mem_set_union @orb_Prop_r @l_req | >nth_opt_append_r [2: >length_reverse @not_lt_to_le assumption ] cases (pt - ?) [ whd in match (nth_opt ???); whd in ⊢ (??%?→?); #EQ destruct(EQ) whd >graph_to_lin_labels in ⊢ (???%); whd in match required'; generalize in match (stmt_explicit_labels … statement); #l @list_as_set_All #i >mem_set_union >mem_set_union #i_l @orb_Prop_r @orb_Prop_l @i_l | @add_req_gen_prf [ #_ | #next #_ #next_vis * [ whd in ⊢ (??%?→?); #EQ' destruct(EQ') whd %{I} >mem_set_union @orb_Prop_l @mem_set_add_id ]] #n whd in ⊢ (??%?→?); #ABS destruct(ABS) ] ] | elim H #pre ** #H1 #H2 #H3 #i whd in match visited'; @(eq_identifier_elim … i vis_hd) [ #EQi >EQi -i #pos >lookup_add_hit #EQpos (* too slow: destruct(EQpos) *) cut (gen_length = pos) [1,3,5: (* BUG: -graph_visit *) -visited destruct(EQpos) %] -EQpos #EQpos lin_code_has_label @add_req_gen_prf [ #_ | #next #_ #next_vis change with (? @ ([?] @ [?])) in match (? @ [? ; ?]); occurs_exactly_once_None ] >occurs_exactly_once_Some_eq >eq_identifier_refl normalize nodelta @generated_prf2 lapply (in_map_domain … visited vis_hd) >H3 normalize nodelta // | %{statement} % [ % ] [ @lookup_eq_safe | nth_opt_append_r >rev_length K #next_vis ] whd >mem_set_add in next_vis; @eq_identifier_elim [1,3: #EQnext >EQnext * [#ABS elim(ABS I)] >lookup_add_hit |*: #NEQ >(lookup_add_miss … visited … NEQ) change with (?∈?) in match (false∨?); #next_vis lapply(in_map_domain … visited next) >next_vis whd in ⊢ (% → ?); [2: * #s ] #EQlookup >EQlookup ] whd [1,2: %2 nth_opt_append_r >append_length >rev_length >commutative_plus K % ] ] ] | #NEQ #n_i >(lookup_add_miss … visited … NEQ) #Hlookup elim (generated_prf1 … Hlookup) #G1 * #s ** #G2 #G3 #G4 % [ >lin_code_has_label occurs_exactly_once_append @orb_Prop_l @andb_Prop [ >occurs_exactly_once_Some_eq >eq_identifier_false [2: % #ABS >ABS in NEQ; * #ABS' @ABS' % ] normalize nodelta >lin_code_has_label in G1; #K @K | @add_req_gen_prf [ #_ % | #next #_ #_ % ] ] | %{s} % [ % ] [ assumption | @nth_opt_append_hit_l assumption | @(opt_All_mp … G4) #x @(eq_identifier_elim … x vis_hd) #Heq [ >Heq lapply (in_map_domain … visited vis_hd) >H3 normalize nodelta in ⊢ (%→?); #EQlookup >EQlookup * #nth_opt_visiting #gen_length_eq >lookup_add_hit %1 >gen_length_eq % | >(lookup_add_miss ?????? Heq) lapply (in_map_domain … visited x) elim (true_or_false_Prop (x∈visited)) #x_vis >x_vis normalize nodelta in ⊢ (%→?); [ * #n' ] #EQlookupx >EQlookupx whd in ⊢ (%→%); * [ #G %1{G} | #G %2 @nth_opt_append_hit_l assumption | #G elim(absurd ?? Heq) (* BUG (but useless): -required -g -generated *) >H2 in G; lapply H1 cases pre [ #_ | #hd #tl * #hd_vis #_ ] normalize #EQ' destruct(EQ') [ % | >x_vis in hd_vis; * ] ] ] ] ] ] | #i whd in match visited'; @(eq_identifier_elim … i vis_hd) #Heq [ >Heq >lookup_add_hit #ABS destruct(ABS) ] >(lookup_add_miss ?????? Heq) #i_n_vis >does_not_occur_append @andb_Prop [ @generated_prf2 assumption | change with (bool_to_Prop (¬eq_identifier ??? ∧ ?)) >eq_identifier_false [2: % #ABS next_vis * | whd in match generated'; @add_req_gen_prf [ #_ | #next #_ #_ ] normalize >gen_length_prf % ] qed. (* CSC: The branch compression (aka tunneling) optimization is not implemented in Matita *) definition branch_compress ≝ λp: graph_params.λglobals.λg:codeT p globals. λentry : Σl.bool_to_Prop (code_has_label … g l).g. lemma branch_compress_closed : ∀p,globals,g,l.code_closed ?? g → code_closed … (branch_compress p globals g l). #p#globals#g#l#H @H qed. lemma branch_compress_has_entry : ∀p,globals,g,l. code_has_label … (branch_compress p globals g l) l. #p#globals#g*#l#l_prf @l_prf qed. definition filter_labels ≝ λtag,A.λtest : identifier tag → bool.λc : list (labelled_obj tag A). map ?? (λs. let 〈l_opt,x〉 ≝ s in 〈! l ← l_opt ; if test l then return l else None ?, x〉) c. lemma does_not_occur_filter_labels : ∀tag,A,test,id,c. does_not_occur ?? id (filter_labels tag A test c) = (does_not_occur ?? id c ∨ ¬ test id). #tag #A #test #id #c elim c [ // | ** [2: #lbl] #s #tl #IH whd in match (filter_labels ????); normalize nodelta whd in match (does_not_occur ????) in ⊢ (??%%); [2: @IH] normalize in match (! l ← ? ; ?); >IH @(eq_identifier_elim ?? lbl id) #Heq [eq_identifier_refl [2: >commutative_orb] normalize % |*: >(eq_identifier_false ??? Heq) normalize nodelta % ] ] qed. lemma occurs_exactly_once_filter_labels : ∀tag,A,test,id,c. occurs_exactly_once ?? id (filter_labels tag A test c) = (occurs_exactly_once ?? id c ∧ test id). #tag #A #test #id #c elim c [ // | ** [2: #lbl] #s #tl #IH whd in match (filter_labels ????); normalize nodelta whd in match (occurs_exactly_once ????) in ⊢ (??%%); [2: @IH] normalize in match (! l ← ? ; ?); >IH >does_not_occur_filter_labels @(eq_identifier_elim ?? lbl id) #Heq [eq_identifier_refl >commutative_andb [ >(commutative_andb ? true) >commutative_orb | >(commutative_andb ? false)] normalize % |*: >(eq_identifier_false ??? Heq) normalize nodelta % ] ] qed. lemma nth_opt_filter_labels : ∀tag,A,test,instrs,n. nth_opt ? n (filter_labels tag A test instrs) = ! 〈lopt, s〉 ← nth_opt ? n instrs ; return 〈 ! lbl ← lopt; if test lbl then return lbl else None ?, s〉. #tag #A #test #instrs elim instrs [ * [2: #n'] % | * #lopt #s #tl #IH * [2: #n'] whd in match (filter_labels ????); normalize nodelta whd in match (nth_opt ???) in ⊢ (??%%); [>IH] % ] qed. lemma stmt_at_filter_labels : ∀p : lin_params.∀globals,test.∀c : codeT p globals. ∀i.stmt_at p globals (filter_labels ?? test c) i = stmt_at p globals c i. #p#globals#test #c#i whd in ⊢ (??%%); >nth_opt_filter_labels elim (nth_opt ???); // qed. lemma option_bind_inverse : ∀A,B.∀m : option A.∀f : A → option B.∀r. ! x ← m ; f x = return r → ∃x.m = return x ∧ f x = return r. #A #B * normalize [2:#x] #f #r #EQ destruct %{x} %{EQ} % qed. lemma nth_opt_reverse_hit : ∀A,l,n.n < |l| → nth_opt A n (reverse ? l) = nth_opt A (|l| - (S n)) l. #A #l elim l [ #n #ABS normalize in ABS; @⊥ -A /2 by absurd/ | #hd #tl #IH #n #lim whd in match (reverse ??); >rev_append_def @(leb_elim (S n) (|tl|)) #H [ >nth_opt_append_l [2: >length_reverse @H ] >(IH … H) >(minus_Sn_m … H) % | >(le_to_le_to_eq … (le_S_S_to_le … lim) (not_lt_to_le … H)) >nth_opt_append_r >length_reverse [2: % ] length_reverse @H qed. definition good_local_sigma : ∀p:unserialized_params. ∀globals. ∀g:codeT (mk_graph_params p) globals. (Σl.bool_to_Prop (code_has_label … g l)) → codeT (mk_lin_params p) globals → (label → option ℕ) → Prop ≝ λp,globals,g,entry,c,sigma. sigma entry = Some ? 0 ∧ ∀l,n.sigma l = Some ? n → ∃s. lookup … g l = Some ? s ∧ opt_Exists ? (λls.let 〈lopt, ts〉 ≝ ls in opt_All ? (eq ? l) lopt ∧ ts = graph_to_lin_statement … s ∧ opt_All ? (λnext.Or (sigma next = Some ? (S n)) (nth_opt … (S n) c = Some ? 〈None ?, GOTO … next〉)) (stmt_implicit_label … s)) (nth_opt … n c). definition linearise_code: ∀p : unserialized_params.∀globals. ∀g : codeT (mk_graph_params p) globals.code_closed … g → ∀entry : (Σl.bool_to_Prop (code_has_label … g l)) .Σ〈c, sigma〉. good_local_sigma … g entry c sigma ∧ code_closed … c (* ∧ ∃ sigma : identifier_map LabelTag ℕ. lookup … sigma entry = Some ? 0 ∧ ∀l,n.lookup … sigma l = Some ? n → ∃s. lookup … g l = Some ? s ∧ opt_Exists ? (λls.let 〈lopt, ts〉 ≝ ls in opt_All ? (eq ? l) lopt ∧ ts = graph_to_lin_statement … s ∧ opt_All ? (λnext.Or (lookup … sigma next = Some ? (S n)) (nth_opt … (S n) c = Some ? 〈None ?, GOTO … next〉)) (stmt_implicit_label … s)) (nth_opt … n c)*) ≝ λp,globals,g,g_prf,entry_sig. let g ≝ branch_compress (mk_graph_params p) ? g entry_sig in let g_prf ≝ branch_compress_closed … g entry_sig g_prf in match graph_visit p globals g ∅ (empty_map …) [ ] [entry_sig] 0 (|g|) (entry_sig) g_prf ????????? with [ mk_Sig triple H ⇒ let sigma ≝ \fst (\fst triple) in let required ≝ \snd (\fst triple) in let crev ≝ \snd triple in let lbld_code ≝ rev ? crev in 〈filter_labels … (λl.l∈required) lbld_code, lookup … sigma〉 ]. [ cases (graph_visit ????????????????????) (* done later *) | #i >mem_set_empty * |3,4: #a #b whd in ⊢ (??%?→?); #ABS destruct(ABS) | #l #_ % | %2 % * >mem_set_empty * | % [2: %] @(branch_compress_has_entry … g entry_sig) | % | % % [% %] cases (pi1 … entry_sig) normalize #_ % // | >commutative_plus change with (? ≤ |g|) % ] ** #visited #required #generated normalize nodelta **** #entry_O #req_vis #last_fin #labels_in_req #sigma_prop % [ % [assumption] #lbl #n #eq_lookup elim (sigma_prop ?? eq_lookup) #lbl_in_gen * #stmt ** #stmt_in_g #nth_opt_is_stmt #succ_is_in % [2: % [ assumption ] |] >nth_opt_filter_labels in ⊢ (???%); >nth_opt_is_stmt >m_return_bind whd >m_return_bind % [ % ] [ elim (lbl ∈ required) % | % | lapply succ_is_in lapply (refl … (stmt_implicit_label … stmt)) cases (stmt_implicit_label … stmt) in ⊢ (???%→%); [#_ #_ %] #next #EQ_next * [ #H %1{H} ] #H %2 >nth_opt_filter_labels >H % ] | #i #s >stmt_at_filter_labels #EQ % [ @stmt_forall_labels_explicit @(All_mp … (labels_in_req … EQ)) #l #l_req >lin_code_has_label >occurs_exactly_once_filter_labels >l_req >commutative_andb whd in ⊢ (?%); elim (sigma_prop ?? (lookup_eq_safe … (req_vis … l_req))) >lin_code_has_label #H #_ @H | lapply EQ cases s #s' [2: #_ % ] * -EQ #EQ change with (bool_to_Prop (code_has_point ????)) whd in match (point_of_succ ???); >lin_code_has_point @leb_elim [ #_ % ] >length_map >length_reverse #INEQ cut (|generated| = S i) [ @(le_to_le_to_eq … (not_lt_to_le … INEQ) ) elim (option_bind_inverse ????? EQ) #x * #EQ1 #EQ2 nth_opt_reverse_hit >EQ_length [2: % ] EQ' #ABS destruct(ABS) ] ] qed. definition linearise_int_fun : ∀p : unserialized_params. ∀globals. ∀fn_in : joint_closed_internal_function (mk_graph_params p) globals .Σ〈fn_out : joint_closed_internal_function (mk_lin_params p) globals, sigma : ?〉. good_local_sigma … (joint_if_code … fn_in) (joint_if_entry … fn_in) (joint_if_code … fn_out) sigma (* ∃sigma : identifier_map LabelTag ℕ. let g ≝ joint_if_code ?? (pi1 … fin) in let c ≝ joint_if_code ?? (pi1 … fout) in let entry ≝ joint_if_entry ?? (pi1 … fin) in lookup … sigma entry = Some ? 0 ∧ ∀l,n.lookup … sigma l = Some ? n → ∃s. lookup … g l = Some ? s ∧ opt_Exists ? (λls.let 〈lopt, ts〉 ≝ ls in opt_All ? (eq ? l) lopt ∧ ts = graph_to_lin_statement … s ∧ opt_All ? (λnext.Or (lookup … sigma next = Some ? (S n)) (nth_opt … (S n) c = Some ? 〈None ?, GOTO … next〉)) (stmt_implicit_label … s)) (nth_opt … n c)*) ≝ λp,globals,f_sig. let code_sigma ≝ linearise_code … (joint_if_code … f_sig) (pi2 … f_sig) (joint_if_entry … f_sig) in let code ≝ \fst code_sigma in let sigma ≝ \snd code_sigma in let entry : Σpt.bool_to_Prop (code_has_point … code pt) ≝ «0, hide_prf ??» in 〈«mk_joint_internal_function (mk_lin_params p) globals (joint_if_luniverse ?? f_sig) (joint_if_runiverse ?? f_sig) (joint_if_result ?? f_sig) (joint_if_params ?? f_sig) (joint_if_locals ?? f_sig) (joint_if_stacksize ?? f_sig) code entry entry (* exit is dummy! *), ?», sigma〉. normalize nodelta cases (linearise_code ?????) * #code #sigma normalize nodelta * #H1 #H2 [ @H2 | @H1 | cases H1 #H3 #H4 elim (H4 … H3) #s * #_ >lin_code_has_point cases code [ * | #hd #tl #_ % ] ] qed. definition linearise : ∀p : unserialized_params. program (joint_function (mk_graph_params p)) ℕ → program (joint_function (mk_lin_params p)) ℕ ≝ λp,pr.transform_program ??? pr (λglobals.transf_fundef ?? (λf_in.\fst (linearise_int_fun p globals f_in))). definition good_sigma : ∀p:unserialized_params. ∀prog_in : joint_program (mk_graph_params p). ((Σi.is_internal_function_of_program … prog_in i) → label → option ℕ) → Prop ≝ λp,prog_in,sigma. let prog_out ≝ linearise … prog_in in ∀i : Σi.is_internal_function_of_program … prog_in i.∀prf. let fn_in ≝ if_of_function … i in let i' : Σi.is_internal_function_of_program … prog_out i ≝ «pi1 ?? i, prf» in let fn_out ≝ if_of_function … i' in let sigma_local ≝ sigma i in good_local_sigma ?? (joint_if_code ?? fn_in) (joint_if_entry … fn_in) (joint_if_code ?? fn_out) sigma_local. lemma linearise_spec : ∀p,prog.∃sigma.good_sigma p prog sigma. #p #prog letin sigma ≝ (λi : Σi.is_internal_function_of_program … prog i. let fn_in ≝ if_of_function … i in \snd (linearise_int_fun … fn_in)) %{sigma} #i #prf >(prog_if_of_function_transform … i) [2: % ] normalize nodelta cases (linearise_int_fun ???) * #fn_out #sigma_loc normalize nodelta #prf @prf qed.