diff --git a/basis/pure/README.md b/basis/pure/README.md index 6572d82b97..b0ccc63db6 100644 --- a/basis/pure/README.md +++ b/basis/pure/README.md @@ -6,6 +6,9 @@ from these by the translator. [basisComputeLib.sml](basisComputeLib.sml): compset for the pure basis functions. +[basis_cvScript.sml](basis_cvScript.sml): +Translation of basis types and functions for use with cv_compute. + [mlintScript.sml](mlintScript.sml): Pure functions for the Int module. diff --git a/basis/pure/basis_cvScript.sml b/basis/pure/basis_cvScript.sml new file mode 100644 index 0000000000..f026254fca --- /dev/null +++ b/basis/pure/basis_cvScript.sml @@ -0,0 +1,53 @@ +(* + Translation of basis types and functions for use with cv_compute. +*) +open preamble cv_transLib cv_stdTheory; + +val _ = new_theory "basis_cv"; + +val _ = cv_memLib.use_long_names := true; + +Triviality list_mem[cv_inline] = listTheory.MEM; + +val _ = cv_trans sptreeTheory.fromAList_def; +val _ = cv_trans miscTheory.SmartAppend_def; +val _ = cv_trans miscTheory.append_aux_def; +val _ = cv_trans miscTheory.append_def; +val _ = cv_trans miscTheory.tlookup_def; +val _ = cv_trans mlstringTheory.implode_def; +val _ = cv_trans mlstringTheory.explode_thm; +val _ = cv_trans mlstringTheory.strcat_thm; +val _ = cv_trans mlstringTheory.str_def; +val _ = cv_auto_trans mlstringTheory.concat_def; +val _ = cv_trans miscTheory.list_max_def; +val _ = cv_trans (miscTheory.max3_def |> PURE_REWRITE_RULE [GREATER_DEF]); + +val toChar_pre = cv_trans_pre mlintTheory.toChar_def +val num_to_chars_pre = cv_auto_trans_pre mlintTheory.num_to_chars_def; + +Theorem IMP_toChar_pre: + n < 16 ⇒ mlint_toChar_pre n +Proof + gvs [toChar_pre] +QED + +Theorem num_to_chars_pre[cv_pre,local]: + ∀a0 a1 a2 a3. mlint_num_to_chars_pre a0 a1 a2 a3 +Proof + ho_match_mp_tac mlintTheory.num_to_chars_ind \\ rw [] + \\ rw [] \\ simp [Once num_to_chars_pre] + \\ once_rewrite_tac [toChar_pre] \\ gvs [] \\ rw [] + \\ ‘k MOD 10 < 10’ by gvs [] \\ simp [] +QED + +Triviality Num_ABS: + Num (ABS i) = Num i +Proof + Cases_on ‘i’ \\ gvs [] +QED + +val _ = cv_trans (mlintTheory.toString_def |> SRULE [Num_ABS]); +val _ = cv_trans mlintTheory.num_to_str_def; + +val _ = Feedback.set_trace "TheoryPP.include_docs" 0; +val _ = export_theory(); diff --git a/compiler/backend/README.md b/compiler/backend/README.md index 26f8545b7e..cfb6729484 100644 --- a/compiler/backend/README.md +++ b/compiler/backend/README.md @@ -1,5 +1,8 @@ The CakeML compiler backend. +[San](San): +A case study for the shared memory feature. + [ag32](ag32): This directory contains the Silver-specific part of the compiler backend and associated proofs. diff --git a/compiler/backend/ag32/proofs/ag32_configProofScript.sml b/compiler/backend/ag32/proofs/ag32_configProofScript.sml index 0493d4bcd8..ce0467e2f6 100644 --- a/compiler/backend/ag32/proofs/ag32_configProofScript.sml +++ b/compiler/backend/ag32/proofs/ag32_configProofScript.sml @@ -79,4 +79,63 @@ val ag32_compile_correct = |> DISCH_ALL |> curry save_thm"ag32_compile_correct"; +Theorem get_shmem_info_LENGTH: + ∀xs c l1 l2. + get_shmem_info xs c l1 l2 = (t1,t2) ∧ LENGTH l1 = LENGTH l2 ⇒ + LENGTH t1 = LENGTH t2 +Proof + ho_match_mp_tac lab_to_targetTheory.get_shmem_info_ind \\ rw [] + \\ gvs [lab_to_targetTheory.get_shmem_info_def] + \\ pairarg_tac \\ gvs [] +QED + +Triviality IMP_EVERY_list_add_if_fresh: + ∀xs x p. p x ∧ EVERY p xs ⇒ EVERY p (list_add_if_fresh x xs) +Proof + Induct \\ gvs [lab_to_targetTheory.list_add_if_fresh_def] \\ rw [] +QED + +Theorem find_ffi_names_ExtCall: + ∀xs. EVERY (λx. ∃i. ExtCall i = x) (find_ffi_names xs) +Proof + ho_match_mp_tac lab_to_targetTheory.find_ffi_names_ind + \\ rpt strip_tac + \\ asm_rewrite_tac [lab_to_targetTheory.find_ffi_names_def,EVERY_DEF] + \\ Cases_on ‘x’ \\ gvs [] + \\ Cases_on ‘a’ \\ gvs [] + \\ irule IMP_EVERY_list_add_if_fresh \\ gvs [] +QED + +Theorem MAP_ExtCall_ffinames: + ∀xs. + EVERY (λx. ∃i. ExtCall i = x) xs ⇒ + xs = MAP ExtCall (ffinames_to_string_list xs) +Proof + Induct \\ gvs [backendTheory.ffinames_to_string_list_def] + \\ Cases \\ gvs [backendTheory.ffinames_to_string_list_def] +QED + +Theorem compile_imp_ffi_names: + backend$compile c p = SOME (b,d,c1) ∧ + c1.lab_conf.shmem_extra = [] ∧ f ≠ [] ∧ + c.lab_conf.ffi_names = NONE ∧ + ffinames_to_string_list (the [] c1.lab_conf.ffi_names) = f ⇒ + c1.lab_conf.ffi_names = SOME (MAP ExtCall f) +Proof + Cases_on ‘c1.lab_conf.ffi_names’ + \\ gvs [libTheory.the_def,backendTheory.ffinames_to_string_list_def] + \\ gvs [backendTheory.compile_def] + \\ rpt (pairarg_tac \\ gvs []) + \\ gvs [oneline backendTheory.attach_bitmaps_def, AllCaseEqs()] + \\ strip_tac \\ gvs [] + \\ gvs [lab_to_targetTheory.compile_def] + \\ gvs [lab_to_targetTheory.compile_lab_def] + \\ rpt (pairarg_tac \\ gvs []) + \\ gvs [AllCaseEqs()] + \\ rpt (pairarg_tac \\ gvs []) + \\ drule get_shmem_info_LENGTH \\ gvs [] \\ rw [] + \\ irule MAP_ExtCall_ffinames + \\ gvs [find_ffi_names_ExtCall] +QED + val _ = export_theory(); diff --git a/compiler/backend/backendScript.sml b/compiler/backend/backendScript.sml index f4d2421fad..44ec9d2878 100644 --- a/compiler/backend/backendScript.sml +++ b/compiler/backend/backendScript.sml @@ -669,7 +669,6 @@ Proof lab_to_targetTheory.shmem_info_num_component_equality]>>fs[] QED - Theorem inc_config_to_config_inv: asm_c = c.lab_conf.asm_conf ==> inc_config_to_config asm_c (config_to_inc_config c) = c @@ -805,4 +804,58 @@ Proof \\ fs [FUN_EQ_THM,FORALL_PROD] QED +Definition ffinames_to_string_list_def: + (ffinames_to_string_list [] = []) ∧ + (ffinames_to_string_list ((ExtCall s)::rest) = + s::(ffinames_to_string_list rest)) ∧ + (ffinames_to_string_list ((SharedMem _)::rest) = + ffinames_to_string_list rest) +End + + +Definition inc_set_oracle_def: + inc_set_oracle c oracle = + c with inc_word_to_word_conf := + (c.inc_word_to_word_conf with col_oracle := oracle) +End + +Definition set_oracle_def: + set_oracle c oracle = + c with + word_to_word_conf := c.word_to_word_conf with col_oracle := oracle +End + +Definition set_asm_conf_def: + set_asm_conf c asm_c = + c with lab_conf := c.lab_conf with asm_conf := asm_c +End + +Theorem inc_set_oracle_pull: + ∀oracle c. inc_config_to_config b (inc_set_oracle c oracle) = + set_oracle (inc_config_to_config b c) oracle +Proof + gvs [inc_set_oracle_def,inc_config_to_config_def,set_oracle_def] +QED + +Theorem inc_config_to_config_config_to_inc_config: + inc_config_to_config asm_c (config_to_inc_config c) = + set_asm_conf c asm_c +Proof + gvs [inc_config_to_config_def, + config_to_inc_config_def, + lab_to_targetTheory.inc_config_to_config_def, + lab_to_targetTheory.config_to_inc_config_def, + config_component_equality, + lab_to_targetTheory.config_component_equality, + set_asm_conf_def] +QED + +Theorem set_asm_conf_id: + c.lab_conf.asm_conf = asm_c ⇒ + set_asm_conf c asm_c = c +Proof + gvs [set_asm_conf_def, fetch "-" "config_component_equality", + lab_to_targetTheory.config_component_equality] +QED + val _ = export_theory(); diff --git a/compiler/backend/bvi_letScript.sml b/compiler/backend/bvi_letScript.sml index 0dd7199088..5c7d3e2f49 100644 --- a/compiler/backend/bvi_letScript.sml +++ b/compiler/backend/bvi_letScript.sml @@ -42,17 +42,20 @@ open preamble bviTheory; val _ = new_theory "bvi_let"; -val extract_def = Define ` +Definition extract_def: (extract ((Var n):bvi$exp) ys = n + LENGTH ys + 1) /\ - (extract _ _ = 0)` + (extract _ _ = 0) +End -val extract_list_def = Define ` +Definition extract_list_def: (extract_list [] = []) /\ - (extract_list (x::xs) = extract x xs :: extract_list xs)` + (extract_list (x::xs) = extract x xs :: extract_list xs) +End -val delete_var_def = Define ` +Definition delete_var_def: (delete_var ((Var n):bvi$exp) = Op (Const 0) []) /\ - (delete_var x = x)`; + (delete_var x = x) +End Theorem exp2_size_APPEND: !xs ys. exp2_size (xs++ys) = exp2_size xs + exp2_size ys @@ -60,7 +63,7 @@ Proof Induct \\ fs [exp_size_def] QED -val compile_def = tDefine "compile" ` +Definition compile_def: (compile env d [] = []) /\ (compile env d (x::y::xs) = compile env d [x] ++ compile env d (y::xs)) /\ (compile env d [(Var v):bvi$exp] = @@ -90,14 +93,66 @@ val compile_def = tDefine "compile" ` (compile env d [Call t dest xs h] = [Call t dest (compile env d xs) (case h of NONE => NONE - | SOME e => SOME (HD (compile (0::env) d [e])))])` - (WF_REL_TAC `measure (bvi$exp2_size o SND o SND)` + | SOME e => SOME (HD (compile (0::env) d [e])))]) +Termination + WF_REL_TAC `measure (bvi$exp2_size o SND o SND)` \\ rw [] \\ fs [LENGTH_NIL] \\ imp_res_tac (METIS_PROVE [SNOC_CASES] ``m <> [] ==> ?x y. m = SNOC x y``) \\ full_simp_tac std_ss [LAST_SNOC,LENGTH_SNOC,FRONT_SNOC] - \\ fs [exp2_size_APPEND,SNOC_APPEND,exp_size_def]); + \\ fs [exp2_size_APPEND,SNOC_APPEND,exp_size_def] +End -val compile_ind = theorem"compile_ind"; +Definition compile_sing_def: + (compile_sing env d ((Var v):bvi$exp) = + case LLOOKUP env v of + | SOME n => Var (v + n) + | _ => Var (v + d)) /\ + (compile_sing env d (If x1 x2 x3) = + (If (compile_sing env d x1) + (compile_sing env d x2) + (compile_sing env d x3))) /\ + (compile_sing env d (Let xs x2) = + let l = LENGTH xs in + if l = 0 then compile_sing env d x2 else + let k = l-1 in + if x2 = Var k then + (* moves the last exp into tail-position if the body is a Var *) + let ys = compile_list env d (BUTLAST xs) in + (Let ys ((compile_sing (MAP ((+)k) env) (d+k) (LAST xs)))) + else + let ys = compile_list env d xs in + (Let (MAP delete_var ys) + ((compile_sing (extract_list ys ++ env) d x2)))) /\ + (compile_sing env d (Raise x1) = + (Raise ((compile_sing env d x1)))) /\ + (compile_sing env d (Op op xs) = (Op op (compile_list env d xs))) /\ + (compile_sing env d (Tick x) = (Tick (compile_sing env d x))) /\ + (compile_sing env d (Call t dest xs h) = + (Call t dest (compile_list env d xs) + (case h of NONE => NONE + | SOME e => SOME ((compile_sing (0::env) d e))))) /\ + (compile_list env d [] = []) /\ + (compile_list env d (x::xs) = + compile_sing env d x :: compile_list env d xs) +Termination + WF_REL_TAC ‘measure $ λx. case x of + | INL (_,_,e) => bvi$exp_size e + | INR (_,_,es) => bvi$exp2_size es’ + \\ conj_tac \\ Cases using SNOC_CASES + \\ gvs [SNOC_APPEND,exp2_size_APPEND,bviTheory.exp_size_def] +End + +Theorem compile_sing: + (∀n d e. compile n d [e] = [compile_sing n d e]) ∧ + (∀n d e. compile_list n d e = compile n d e) +Proof + ho_match_mp_tac compile_sing_ind \\ rpt strip_tac + \\ gvs [compile_sing_def,compile_def] + \\ every_case_tac \\ gvs [] + \\ rename [‘compile n d (e::es)’] + \\ Cases_on ‘es’ \\ gvs [] + \\ gvs [compile_sing_def,compile_def] +QED Theorem compile_length[simp]: !n d xs. LENGTH (compile n d xs) = LENGTH xs @@ -114,7 +169,11 @@ Proof \\ Cases_on `compile n d [x]` \\ fs [LENGTH_NIL] QED -val compile_exp_def = Define ` - compile_exp x = case compile [] 0 [x] of (y::_) => y | _ => Var 0 (* impossible *)`; +Definition compile_exp_def: + compile_exp x = + case compile [] 0 [x] of (y::_) => y | _ => Var 0 (* impossible *) +End + +Theorem compile_exp_eq = compile_exp_def |> SRULE [compile_sing]; val _ = export_theory(); diff --git a/compiler/backend/bvi_tailrecScript.sml b/compiler/backend/bvi_tailrecScript.sml index 44b40aa8a0..39f3b2c543 100644 --- a/compiler/backend/bvi_tailrecScript.sml +++ b/compiler/backend/bvi_tailrecScript.sml @@ -528,6 +528,77 @@ val scan_expr_def = tDefine "scan_expr" ` [(ts, Any, F, NONE)])` (WF_REL_TAC `measure (exp2_size o SND o SND)`); + +Definition scan_expr_sing_def: + (scan_expr_sing ts loc (Var n) = + let ty = if n < LENGTH ts then EL n ts else Any in + (ts, ty, F, NONE)) ∧ + (scan_expr_sing ts loc (If xi xt xe) = + let (ti, tyi, _, oi) = scan_expr_sing ts loc xi in + let (tt, ty1, _, ot) = scan_expr_sing ti loc xt in + let (te, ty2, _, oe) = scan_expr_sing ti loc xe in + let op = dtcase ot of NONE => oe | _ => ot in + let ty = dtcase op of + NONE => decide_ty ty1 ty2 + | SOME opr => decide_ty ty1 (decide_ty ty2 (op_type opr)) in + (MAP2 decide_ty tt te, ty, IS_SOME oe, op)) ∧ + (scan_expr_sing ts loc (Let xs x) = + let ys = scan_expr_list ts loc xs in + let tt = MAP (FST o SND) ys in + let tr = (dtcase LAST1 ys of SOME c => FST c | NONE => ts) in + let (tu, ty, _, op) = scan_expr_sing (tt ++ tr) loc x in + (DROP (LENGTH ys) tu, ty, F, op)) ∧ + (scan_expr_sing ts loc (Raise x) = (ts, Any, F, NONE)) ∧ + (scan_expr_sing ts loc (Tick x) = scan_expr_sing ts loc x) ∧ + (scan_expr_sing ts loc (Call t d xs h) = (ts, Any, F, NONE)) ∧ + (scan_expr_sing ts loc (Op op xs) = + let opr = from_op op in + let opt = op_type opr in + dtcase opr of + Noop => (* Constants? *) + if arg_ty op = Int then + dtcase get_bin_args (Op op xs) of + NONE => + if is_const op then + (ts, Int, F, NONE) + else + (ts, Any, F, NONE) + | SOME (x, y) => + if ~is_const op then + (update_context Int ts x y, Any, F, NONE) + else + (ts, Any, F, NONE) + else if op = Cons 0 /\ xs = [] then (* list nil *) + (ts, List, F, NONE) + else + (ts, Any, F, NONE) + | _ => (* Things we can optimize *) + if IS_SOME (check_op ts opr loc (Op op xs)) then + (ts, opt, F, SOME opr) + else if term_ok ts opt (Op op xs) then + dtcase get_bin_args (Op op xs) of + NONE => (ts, Any, F, NONE) + | SOME (x, y) => + (update_context opt ts x y, opt, F, NONE) + else + (ts, Any, F, NONE)) ∧ + + (scan_expr_list ts loc [] = []) ∧ + (scan_expr_list ts loc (x::xs) = + let (tx, ty, r, op) = scan_expr_sing ts loc x in + (tx, ty, r, op)::scan_expr_list tx loc xs) +End + +Theorem scan_expr_eq: + (∀e ts loc. scan_expr ts loc [e] = [scan_expr_sing ts loc e]) ∧ + (∀es ts loc. scan_expr ts loc es = scan_expr_list ts loc es) +Proof + Induct >> rw[] >> gvs[scan_expr_def, scan_expr_sing_def] >> + rpt (pairarg_tac >> gvs[]) >> + rpt (TOP_CASE_TAC >> gvs[]) >> + Cases_on `es` >> simp[Once scan_expr_def, scan_expr_sing_def] +QED + val push_call_def = Define ` (push_call n op acc exp (SOME (ticks, dest, args, handler)) = Call ticks (SOME n) (args ++ [apply_op op (Var acc) exp]) handler) ∧ @@ -592,6 +663,10 @@ Proof \\ recInduct (theorem "rewrite_ind") \\ rw [rewrite_def] QED + +Theorem rewrite_eq = rewrite_def |> SRULE [scan_expr_eq]; + + (* --- Top-level expression check --- *) val has_rec_def = tDefine "has_rec" ` @@ -612,6 +687,35 @@ val has_rec_def = tDefine "has_rec" ` val has_rec1_def = Define `has_rec1 loc x = has_rec loc [x]`; +Definition has_rec_sing_def: + (has_rec_sing loc (If x1 x2 x3) = + if has_rec_sing loc x2 then T + else has_rec_sing loc x3) /\ + (has_rec_sing loc (Let xs x) = has_rec_sing loc x) /\ + (has_rec_sing loc (Tick x) = has_rec_sing loc x) /\ + (has_rec_sing loc (Op op xs) = + if EXISTS (is_rec loc) xs then T + else has_rec_list loc xs) /\ + (has_rec_sing loc x = F) ∧ + + (has_rec_list loc [] = F) /\ + (has_rec_list loc (x::xs) = + if has_rec_sing loc x then T + else has_rec_list loc xs) +Termination + WF_REL_TAC `measure (λx. case x of INL (_,e) => exp_size e + | INR (_,es) => list_size exp_size es)` >> + simp[bviTheory.exp_size_def, listTheory.list_size_def, bviTheory.exp_size_eq] +End + +Theorem has_rec_eq: + (∀e loc. has_rec loc [e] = has_rec_sing loc e) ∧ + (∀es loc. has_rec loc es = has_rec_list loc es) +Proof + Induct >> rw[has_rec_def, has_rec_sing_def] >> + Cases_on `es` >> gvs[has_rec_def] +QED + val test1_tm = ``Let [] (Call 0 (SOME 0) [] NONE)`` Theorem has_rec_test1: has_rec1 0 ^test1_tm <=> F @@ -638,6 +742,9 @@ val check_exp_def = Define ` | SOME op => if ty <> op_type op then NONE else opr`; +Theorem check_exp_eq = check_exp_def |> + SRULE [scan_expr_eq, has_rec1_def, has_rec_eq]; + val let_wrap_def = Define ` let_wrap arity id exp = Let ((GENLIST (λi. Var i) arity) ++ [id]) exp`; diff --git a/compiler/backend/bvi_to_dataScript.sml b/compiler/backend/bvi_to_dataScript.sml index cfb11a0add..14bf4a9b4a 100644 --- a/compiler/backend/bvi_to_dataScript.sml +++ b/compiler/backend/bvi_to_dataScript.sml @@ -82,7 +82,7 @@ val iAssign_def = Define ` val _ = Parse.hide"tail"; -val compile_def = tDefine "compile" ` +Definition compile_def: (compile (n:num) (env:num list) tail live [] = (Skip,[]:num list,n)) /\ (compile n env tail live (x::y::xs) = @@ -127,10 +127,85 @@ val compile_def = tDefine "compile" ` let (c2,v,n2) = compile (n1+1) (n1::env) F live [handler] in let ret = SOME (n2, list_to_num_set (live ++ env)) in let c3 = (if tail then Return n2 else Skip) in - (Seq c1 (mk_ticks ticks (Seq (Call ret dest vs (SOME (n1,Seq c2 (Move n2 (HD v))))) c3)), [n2], n2+1))` - (WF_REL_TAC `measure (exp2_size o SND o SND o SND o SND)`); + (Seq c1 (mk_ticks ticks (Seq (Call ret dest vs (SOME (n1,Seq c2 (Move n2 (HD v))))) c3)), [n2], n2+1)) +Termination + WF_REL_TAC `measure (exp2_size o SND o SND o SND o SND)` +End -val compile_ind = theorem"compile_ind"; +Definition compile_sing_def: + (compile_sing n env tail live ((Var v):bvi$exp) = + let var = any_el v env 0n in + if tail + then (Return var, n, n+1) + else (Skip, var, MAX n (var+1))) /\ + (compile_sing n env tail live (If x1 x2 x3) = + let (c1,v1,n1) = compile_sing n env F live x1 in + let (c2,v2,n2) = compile_sing n1 env tail live x2 in + let (c3,v3,n3) = compile_sing n2 env tail live x3 in + if tail then + (Seq c1 (If v1 c2 c3),n3,n3+1) + else + (Seq c1 (If v1 (Seq c2 (Move n3 v2)) + (Seq c3 (Move n3 v3))),n3,n3+1)) /\ + (compile_sing n env tail live (Let xs x2) = + let (c1,vs,n1) = compile_list n env live xs in + let (c2,v2,n2) = compile_sing n1 (vs ++ env) tail live x2 in + (Seq c1 c2, v2, n2)) /\ + (compile_sing n env tail live (Raise x1) = + let (c1,v1,n1) = compile_sing n env F live x1 in + (Seq c1 (Raise v1), v1, n1)) /\ + (compile_sing n env tail live (Op op xs) = + let (c1,vs,n1) = compile_list n env live xs in + let c2 = Seq c1 (iAssign n1 op (REVERSE vs) live env) in + (if tail then Seq c2 (Return n1) else c2, n1, n1+1)) /\ + (compile_sing n env tail live (Tick x1) = + let (c1,v1,n1) = compile_sing n env tail live x1 in + (Seq Tick c1, v1, n1)) /\ + (compile_sing n env tail live (Call ticks dest xs NONE) = + let (c1,vs,n1) = compile_list n env live xs in + let ret = (if tail then NONE + else SOME (n1, list_to_num_set (live ++ env))) in + (Seq c1 (mk_ticks ticks (Call ret dest vs NONE)), n1, n1+1)) /\ + (compile_sing n env tail live (Call ticks dest xs (SOME handler)) = + let (c1,vs,n1) = compile_list n env live xs in + let (c2,v,n2) = compile_sing (n1+1) (n1::env) F live handler in + let ret = SOME (n2, list_to_num_set (live ++ env)) in + let c3 = (if tail then Return n2 else Skip) in + (Seq c1 (mk_ticks ticks + (Seq (Call ret dest vs (SOME (n1,Seq c2 (Move n2 v)))) c3)), + n2, n2+1)) /\ + (compile_list (n:num) (env:num list) live [] = + (Skip,[]:num list,n)) /\ + (compile_list n env live (x::xs) = + let (c1,v1,n1) = compile_sing n env F live x in + if NULL xs then (c1,[v1],n1) else + let (c2,vs,n2) = compile_list n1 env (v1::live) xs in + (Seq c1 c2, v1 :: vs, n2)) +Termination + WF_REL_TAC ‘measure $ λx. case x of + | INL (n,env,t,l,x) => exp_size x + | INR (n,env,l,xs) => exp2_size xs’ + \\ rw [] \\ gvs [exp_size_def] +End + +Theorem compile_sing_eq: + (∀n env t l x c1 v1 n1. + compile_sing n env t l x = (c1,v1,n1) ⇒ + compile n env t l [x] = (c1,[v1],n1)) ∧ + (∀n env l xs. + compile_list n env l xs = + compile n env F l xs) +Proof + ho_match_mp_tac compile_sing_ind \\ rw [] + \\ once_rewrite_tac [compile_sing_def,compile_def] \\ gvs [] + \\ rpt (pairarg_tac \\ gvs []) + \\ gvs [compile_sing_def] + \\ rpt (IF_CASES_TAC \\ gvs []) + \\ rpt (pairarg_tac \\ gvs []) + \\ Cases_on ‘xs’ \\ gvs [compile_sing_def] + \\ rpt (pairarg_tac \\ gvs []) + \\ gvs [compile_def] +QED val compile_LESS_EQ_lemma = Q.prove( `!n env tail live xs. @@ -178,15 +253,28 @@ val optimise_def = Define ` (* the top-level compiler includes the optimisations, because the correctness proofs are combined *) -val compile_exp = Define ` +Definition compile_exp_def: compile_exp arg_count exp = - optimise (FST (compile arg_count (COUNT_LIST arg_count) T [] [exp]))` + optimise (FST (compile arg_count (COUNT_LIST arg_count) T [] [exp])) +End -val compile_part = Define ` +Definition compile_part_def: compile_part (name:num, arg_count, exp) = - (name, arg_count, compile_exp arg_count exp)` + (name, arg_count, compile_exp arg_count exp) +End -val compile_prog_def = Define ` - compile_prog prog = MAP compile_part prog`; +Definition compile_prog_def: + compile_prog prog = MAP compile_part prog +End + +Theorem compile_exp_eq: + compile_exp arg_count exp = + optimise (FST (compile_sing arg_count (COUNT_LIST arg_count) T [] exp)) +Proof + ‘∃z. compile_sing arg_count (COUNT_LIST arg_count) T [] exp = z’ by fs [] + \\ PairCases_on ‘z’ \\ gvs [] + \\ drule $ cj 1 compile_sing_eq + \\ gvs [compile_exp_def] +QED val _ = export_theory(); diff --git a/compiler/backend/bvl_constScript.sml b/compiler/backend/bvl_constScript.sml index 1bfcecc3dd..bb14b35d36 100644 --- a/compiler/backend/bvl_constScript.sml +++ b/compiler/backend/bvl_constScript.sml @@ -317,6 +317,45 @@ Termination WF_REL_TAC `measure (exp1_size o SND)` End +Definition compile_sing_def: + (compile_sing env (Var v) = + dtcase LLOOKUP env v of + | NONE => Var v + | SOME NONE => Var v + | SOME (SOME (Var i)) => Var (v + i) + | SOME (SOME x) => x) /\ + (compile_sing env (If x1 x2 x3) = + let y1 = compile_sing env x1 in + let y2 = compile_sing env x2 in + let y3 = compile_sing env x3 in + if y1 = Bool T then y2 else + if y1 = Bool F then y3 else + If y1 y2 y3) /\ + (compile_sing env (Let xs x2) = + let ys = compile_list env xs in + Let (MAP delete_var ys) + (compile_sing (extract_list ys ++ env) x2)) /\ + (compile_sing env (Handle x1 x2) = + Handle (compile_sing env x1) (compile_sing (NONE::env) x2)) /\ + (compile_sing env (Raise x1) = + Raise (compile_sing env x1)) /\ + (compile_sing env (Op op xs) = SmartOp op (compile_list env xs)) /\ + (compile_sing env (Tick x) = Tick (compile_sing env x)) /\ + (compile_sing env (Call t dest xs) = Call t dest (compile_list env xs)) ∧ + + (compile_list env [] = []) /\ + (compile_list env (x::xs) = compile_sing env x :: compile_list env xs) +End + +Theorem compile_eq: + (∀e env. compile env [e] = [compile_sing env e]) ∧ + (∀es env. compile env es = compile_list env es) +Proof + Induct >> rw[compile_def, compile_sing_def] >> + rpt (TOP_CASE_TAC >> gvs[]) >> + Cases_on `es` >> simp[compile_def, compile_sing_def] +QED + Theorem compile_length[simp]: !n xs. LENGTH (compile n xs) = LENGTH xs Proof @@ -337,4 +376,6 @@ Definition compile_exp_def: dtcase compile [] [x] of (y::_) => y | _ => Var 0 (* impossible *) End +Theorem compile_exp_eq = compile_exp_def |> SRULE [compile_eq]; + val _ = export_theory(); diff --git a/compiler/backend/bvl_handleScript.sml b/compiler/backend/bvl_handleScript.sml index b253ad86c6..3b585195a1 100644 --- a/compiler/backend/bvl_handleScript.sml +++ b/compiler/backend/bvl_handleScript.sml @@ -136,6 +136,21 @@ val OptionalLetLet_def = Define ` let flat_live = vars_from_list (MAP FST (toAList fvs)) in ([(Let [] (LetLet env_size fvs e))],flat_live,0n,nr)`; +Definition OptionalLetLet_sing_def: + OptionalLetLet_sing e env_size live s (limit:num) (nr:bool) = + if s < limit then (e,live,s,nr) else + let fvs = db_to_set live in + let flat_live = vars_from_list (MAP FST (toAList fvs)) in + ((Let [] (LetLet env_size fvs e)),flat_live,0n,nr) +End + +Theorem OptionalLetLet_sing_eq: + OptionalLetLet e sz lv s lim nr = + (λ(d,l,s,nr). ([d],l,s,nr)) (OptionalLetLet_sing e sz lv s lim nr) +Proof + rw[OptionalLetLet_def, OptionalLetLet_sing_def] +QED + val compile_def = tDefine "compile" ` (compile l n [] = ([]:bvl$exp list,Empty,0n,T)) /\ (compile l n (x::y::xs) = @@ -184,10 +199,74 @@ val compile_def = tDefine "compile" ` val compile_ind = theorem"compile_ind"; +Definition compile_sing_def: + (compile_sing l n (Var v) = if v < n then (Var v,Var v,1,T) + else (Op (Const 0) [],Empty,1,T)) /\ + (compile_sing l n (If x1 x2 x3) = + let (x1,l1,s1,nr1) = compile_sing l n x1 in + let (x2,l2,s2,nr2) = compile_sing l n x2 in + let (x3,l3,s3,nr3) = compile_sing l n x3 in + OptionalLetLet_sing (If x1 x2 x3) n + (mk_Union l1 (mk_Union l2 l3)) (s1+s2+s3+1) l + (nr1 /\ nr2 /\ nr3)) /\ + (compile_sing l n (Let xs x2) = + if NULL xs then + compile_sing l n x2 + else + let k = LENGTH xs in + let (xs,l1,s1,nr1) = compile_list l n xs in + let (x2,l2,s2,nr2) = compile_sing l (n + k) x2 in + OptionalLetLet_sing (Let xs x2) n + (mk_Union l1 (Shift k l2)) (s1+s2+1) l (nr1 /\ nr2)) /\ + (compile_sing l n (Handle x1 x2) = + let (y1,l1,s1,nr1) = compile_sing l n x1 in + if nr1 then (y1,l1,s1,T) else + let (y2,l2,s2,nr2) = compile_sing l (n+1) x2 in + (Handle (LetLet n (db_to_set l1) y1) y2, + mk_Union l1 (Shift 1 l2), + s2 (* s1 intentionally left out because + it does not contrib to exp size in BVI *), nr2)) /\ + (compile_sing l n (Raise x1) = + let (dx,lx,s1,nr1) = compile_sing l n x1 in + OptionalLetLet_sing (Raise dx) n lx (s1+1) l F) /\ + (compile_sing l n (Op op xs) = + let (ys,lx,s1,nr1) = compile_list l n xs in + OptionalLetLet_sing (Op op ys) n lx (s1+1) l nr1) /\ + (compile_sing l n (Tick x) = + let (y,lx,s1,nr1) = compile_sing l n x in + (Tick y,lx,s1,nr1)) /\ + (compile_sing l n (Call t dest xs) = + let (ys,lx,s1,nr1) = compile_list l n xs in + OptionalLetLet_sing (Call t dest ys) n lx (s1+1) l F) ∧ + + (compile_list l n [] = ([]:bvl$exp list,Empty,0n,T)) /\ + (compile_list l n (x::xs) = + let (dx,lx,s1,nr1) = compile_sing l n x in + let (dy,ly,s2,nr2) = compile_list l n xs in + (dx :: dy, mk_Union lx ly, s1+s2, nr1 /\ nr2)) +End + +Theorem compile_sing_eq: + (∀e l n. compile l n [e] = (λ(d,l,s,nr). ([d],l,s,nr)) (compile_sing l n e)) ∧ + (∀es l n. compile l n es = compile_list l n es) +Proof + Induct >> rw[compile_def, compile_sing_def, OptionalLetLet_sing_eq] >> + rpt (pairarg_tac >> gvs[]) >> + rpt (TOP_CASE_TAC >> gvs[]) >> + Cases_on `es` >> gvs[compile_def] +QED + val compile_exp_def = Define ` compile_exp cut_size arity e = HD (FST (compile cut_size arity [handle_simp (bvl_const$compile_exp e)]))`; +Theorem compile_exp_eq: + compile_exp cut_size arity e = + FST (compile_sing cut_size arity (handle_simp (bvl_const$compile_exp e))) +Proof + rw[compile_exp_def, compile_sing_eq] >> pairarg_tac >> gvs[] +QED + val dest_Seq_def = PmatchHeuristics.with_classic_heuristic Define ` (dest_Seq (Let [e1;e2] (Var 1)) = SOME (e1,e2)) /\ (dest_Seq _ = NONE)` diff --git a/compiler/backend/bvl_inlineScript.sml b/compiler/backend/bvl_inlineScript.sml index bc91516723..85d23386e2 100644 --- a/compiler/backend/bvl_inlineScript.sml +++ b/compiler/backend/bvl_inlineScript.sml @@ -9,7 +9,7 @@ val _ = new_theory "bvl_inline"; (* tick_inline -- a function that inlines a function body *) -val tick_inline_def = tDefine "tick_inline" ` +Definition tick_inline_def: (tick_inline cs [] = []) /\ (tick_inline cs (x::y::xs) = HD (tick_inline cs [x]) :: tick_inline cs (y::xs)) /\ @@ -25,7 +25,7 @@ val tick_inline_def = tDefine "tick_inline" ` [Raise (HD (tick_inline cs [x1]))]) /\ (tick_inline cs [Handle x1 x2] = [Handle (HD (tick_inline cs [x1])) - (HD (tick_inline cs [x2]))]) /\ + (HD (tick_inline cs [x2]))]) /\ (tick_inline cs [Op op xs] = [Op op (tick_inline cs xs)]) /\ (tick_inline cs [Tick x] = @@ -34,13 +34,51 @@ val tick_inline_def = tDefine "tick_inline" ` case dest of NONE => [Call ticks dest (tick_inline cs xs)] | SOME n => case lookup n cs of | NONE => [Call ticks dest (tick_inline cs xs)] - | SOME (arity,code) => [Let (tick_inline cs xs) (mk_tick (SUC ticks) code)])` - (WF_REL_TAC `measure (exp1_size o SND)`); + | SOME (arity,code) => [Let (tick_inline cs xs) (mk_tick (SUC ticks) code)]) +Termination + WF_REL_TAC `measure (exp1_size o SND)` +End -val tick_inline_ind = theorem"tick_inline_ind"; +Definition tick_inline_sing_def: + (tick_inline_sing cs (Var v) = Var v) /\ + (tick_inline_sing cs (If x1 x2 x3) = + (If (tick_inline_sing cs x1) + (tick_inline_sing cs x2) + (tick_inline_sing cs x3))) /\ + (tick_inline_sing cs (Let xs x2) = + (Let (tick_inline_list cs xs) + (tick_inline_sing cs x2))) /\ + (tick_inline_sing cs (Raise x1) = + (Raise (tick_inline_sing cs x1))) /\ + (tick_inline_sing cs (Handle x1 x2) = + (Handle (tick_inline_sing cs x1) + (tick_inline_sing cs x2))) /\ + (tick_inline_sing cs (Op op xs) = + (Op op (tick_inline_list cs xs))) /\ + (tick_inline_sing cs (Tick x) = + (Tick (tick_inline_sing cs x))) /\ + (tick_inline_sing cs (Call ticks dest xs) = + case dest of NONE => Call ticks dest (tick_inline_list cs xs) | SOME n => + case lookup n cs of + | NONE => Call ticks dest (tick_inline_list cs xs) + | SOME (arity,code) => + Let (tick_inline_list cs xs) (mk_tick (SUC ticks) code)) ∧ + (tick_inline_list cs [] = []) /\ + (tick_inline_list cs (x::xs) = + (tick_inline_sing cs x) :: tick_inline_list cs xs) +End + +Theorem tick_inline_sing: + (∀e. tick_inline cs [e] = [tick_inline_sing cs e]) ∧ + (∀es. tick_inline_list cs es = tick_inline cs es) +Proof + Induct \\ gvs [tick_inline_sing_def,tick_inline_def] + \\ rw [] \\ every_case_tac + \\ Cases_on ‘es’ \\ gvs [tick_inline_sing_def,tick_inline_def] +QED (* This definition is written to exit as soon as possible. *) -val is_small_aux_def = tDefine "is_small_aux" ` +Definition is_small_aux_def: (is_small_aux n [] = n) /\ (is_small_aux n (x::y::xs) = if n = 0n then n else @@ -69,13 +107,61 @@ val is_small_aux_def = tDefine "is_small_aux" ` (is_small_aux n [Tick x] = is_small_aux n [x]) /\ (is_small_aux n [Call ticks dest xs] = let n = n - 1 in if n = 0 then 0 else - is_small_aux n xs)` - (WF_REL_TAC `measure (exp1_size o SND)`); + is_small_aux n xs) +Termination + WF_REL_TAC `measure (exp1_size o SND)` +End -val is_small_def = Define ` - is_small limit e = ~(is_small_aux limit [e] = 0)`; +Definition is_small_sing_def: + (is_small_sing n (Var v) = n) /\ + (is_small_sing n (If x1 x2 x3) = + let n = n - 1 in if n = 0 then 0 else + let n = is_small_sing n x1 in if n = 0 then 0 else + let n = is_small_sing n x2 in if n = 0 then 0 else + is_small_sing n x3) /\ + (is_small_sing n (Let xs x2) = + let n = n - 1 in if n = 0 then 0 else + let n = is_small_list n xs in if n = 0 then 0 else + is_small_sing n x2) /\ + (is_small_sing n (Raise x1) = + let n = n - 1 in if n = 0 then 0 else + is_small_sing n x1) /\ + (is_small_sing n (Handle x1 x2) = + let n = n - 1 in if n = 0 then 0 else + let n = is_small_sing n x1 in if n = 0 then 0 else + is_small_sing n x2) /\ + (is_small_sing n (Op op xs) = + let n = n - 1 in if n = 0 then 0 else + is_small_list n xs) /\ + (is_small_sing n (Tick x) = is_small_sing n x) /\ + (is_small_sing n (Call ticks dest xs) = + let n = n - 1 in if n = 0 then 0 else + is_small_list n xs) /\ + (is_small_list n [] = n) /\ + (is_small_list n (x::xs) = + if n = 0n then n else + let n = is_small_sing n x in if n = 0 then n else + is_small_list n xs) +End + +Theorem is_small_sing: + (∀e n. is_small_aux n [e] = is_small_sing n e) ∧ + (∀es n. is_small_list n es = is_small_aux n es) +Proof + Induct \\ gvs [is_small_sing_def,is_small_aux_def] \\ rw [] + \\ Cases_on ‘es’ \\ gvs [is_small_sing_def,is_small_aux_def] + \\ qsuff_tac ‘(∀e. is_small_sing 0 e = 0) ∧ + (∀e. is_small_list 0 e = 0)’ \\ gvs [] + \\ Induct \\ gvs [is_small_sing_def] +QED + +Definition is_small_def: + is_small limit e = ~(is_small_aux limit [e] = 0) +End -val is_rec_def = tDefine "is_rec" ` +Theorem is_small_eq = is_small_def |> SRULE [is_small_sing]; + +Definition is_rec_def: (is_rec n [] = F) /\ (is_rec n (x::y::xs) = if is_rec n [x] then T else @@ -95,22 +181,63 @@ val is_rec_def = tDefine "is_rec" ` (is_rec n [Op op xs] = is_rec n xs) /\ (is_rec n [Tick x] = is_rec n [x]) /\ (is_rec n [Call ticks dest xs] = - if dest = SOME n then T else is_rec n xs)` - (WF_REL_TAC `measure (exp1_size o SND)`); + if dest = SOME n then T else is_rec n xs) +Termination + WF_REL_TAC `measure (exp1_size o SND)` +End + +Definition is_rec_sing_def: + (is_rec_sing n (Var v) = F) /\ + (is_rec_sing n (If x1 x2 x3) = + if is_rec_sing n x1 then T else + if is_rec_sing n x2 then T else + is_rec_sing n x3) /\ + (is_rec_sing n (Let xs x2) = + if is_rec_list n xs then T else + is_rec_sing n x2) /\ + (is_rec_sing n (Raise x1) = is_rec_sing n x1) /\ + (is_rec_sing n (Handle x1 x2) = + if is_rec_sing n x1 then T else + is_rec_sing n x2) /\ + (is_rec_sing n (Op op xs) = is_rec_list n xs) /\ + (is_rec_sing n (Tick x) = is_rec_sing n x) /\ + (is_rec_sing n (Call ticks dest xs) = + if dest = SOME n then T else is_rec_list n xs) /\ + (is_rec_list n [] = F) /\ + (is_rec_list n (x::xs) = + if is_rec_sing n x then T else + is_rec_list n xs) +End -val must_inline_def = Define ` +Theorem is_rec_sing: + (∀e n. is_rec n [e] = is_rec_sing n e) ∧ + (∀es n. is_rec_list n es = is_rec n es) +Proof + Induct \\ gvs [is_rec_sing_def,is_rec_def] \\ rw [] + \\ Cases_on ‘es’ \\ gvs [is_rec_sing_def,is_rec_def] +QED + +Definition must_inline_def: must_inline name limit e = - if is_small limit e then ~(is_rec name [e]) else F` + if is_small limit e then ~(is_rec name [e]) else F +End + +Theorem must_inline_eq = must_inline_def |> SRULE [is_rec_sing]; -val tick_inline_all_def = Define ` +Definition tick_inline_all_def: (tick_inline_all limit cs [] aux = (cs,REVERSE aux)) /\ (tick_inline_all limit cs ((n,arity:num,e1)::xs) aux = let e2 = HD (tick_inline cs [e1]) in let cs2 = if must_inline n limit e2 then insert n (arity,e2) cs else cs in - tick_inline_all limit cs2 xs ((n,arity,e2)::aux))`; + tick_inline_all limit cs2 xs ((n,arity,e2)::aux)) +End -val tick_compile_prog_def = Define ` - tick_compile_prog limit cs prog = tick_inline_all limit cs prog []` +Theorem tick_inline_all_eq = + tick_inline_all_def |> SRULE [tick_inline_sing]; + +Definition tick_compile_prog_def: + tick_compile_prog limit cs prog = tick_inline_all limit cs prog [] +End Theorem LENGTH_tick_inline: !cs xs. LENGTH (tick_inline cs xs) = LENGTH xs @@ -129,7 +256,7 @@ QED (* remove_ticks -- a function that removes Ticks *) -val remove_ticks_def = tDefine "remove_ticks" ` +Definition remove_ticks_def: (remove_ticks [] = []) /\ (remove_ticks (x::y::xs) = HD (remove_ticks [x]) :: remove_ticks (y::xs)) /\ @@ -149,8 +276,41 @@ val remove_ticks_def = tDefine "remove_ticks" ` [Op op (remove_ticks xs)]) /\ (remove_ticks [Tick x] = remove_ticks [x]) /\ (remove_ticks [Call ticks dest xs] = - [Call 0 dest (remove_ticks xs)])` - (WF_REL_TAC `measure exp1_size`); + [Call 0 dest (remove_ticks xs)]) +Termination + WF_REL_TAC `measure exp1_size` +End + +Definition remove_ticks_sing_def: + (remove_ticks_sing (Var v) = Var v) /\ + (remove_ticks_sing (If x1 x2 x3) = + If (remove_ticks_sing x1) + (remove_ticks_sing x2) + (remove_ticks_sing x3)) /\ + (remove_ticks_sing (Let xs x2) = + Let (remove_ticks_list xs) (remove_ticks_sing x2)) /\ + (remove_ticks_sing (Raise x1) = + Raise (remove_ticks_sing x1)) /\ + (remove_ticks_sing (Handle x1 x2) = + Handle (remove_ticks_sing x1) + (remove_ticks_sing x2)) /\ + (remove_ticks_sing (Op op xs) = + Op op (remove_ticks_list xs)) /\ + (remove_ticks_sing (Tick x) = remove_ticks_sing x) /\ + (remove_ticks_sing (Call ticks dest xs) = + Call 0 dest (remove_ticks_list xs)) /\ + (remove_ticks_list [] = []) /\ + (remove_ticks_list (x::xs) = + (remove_ticks_sing x) :: remove_ticks_list xs) +End + +Theorem remove_ticks_sing: + (∀e. remove_ticks [e] = [remove_ticks_sing e]) ∧ + (∀es. remove_ticks_list es = remove_ticks es) +Proof + Induct \\ gvs [remove_ticks_sing_def,remove_ticks_def] \\ rw [] + \\ Cases_on ‘es’ \\ gvs [remove_ticks_sing_def,remove_ticks_def] +QED Theorem LENGTH_remove_ticks[simp]: !xs. LENGTH (remove_ticks xs) = LENGTH xs @@ -168,16 +328,18 @@ QED (* let_op -- a function that optimises Let [...] (Op op [Var ...]) *) -val var_list_def = Define ` +Definition var_list_def: var_list n [] [] = T /\ var_list n (bvl$Var m :: xs) (y::ys) = (m = n /\ var_list (n+1) xs ys) /\ - var_list _ _ _ = F` + var_list _ _ _ = F +End -val dest_op_def = Define ` +Definition dest_op_def: dest_op (bvl$Op op xs) args = (if var_list 0 xs args then SOME op else NONE) /\ - dest_op _ _ = NONE` + dest_op _ _ = NONE +End -val let_op_def = tDefine "let_op" ` +Definition let_op_def: (let_op [] = []) /\ (let_op ((x:bvl$exp)::y::xs) = HD (let_op [x]) :: let_op (y::xs)) /\ @@ -200,27 +362,72 @@ val let_op_def = tDefine "let_op" ` (let_op [Op op xs] = [Op op (let_op xs)]) /\ (let_op [Tick x] = [Tick (HD (let_op [x]))]) /\ - (let_op [Call ticks dest xs] = [Call ticks dest (let_op xs)])` - (WF_REL_TAC `measure exp1_size`); + (let_op [Call ticks dest xs] = [Call ticks dest (let_op xs)]) +Termination + WF_REL_TAC `measure exp1_size` +End + +Definition let_op_one_def: + (let_op_one (Var v) = Var v) /\ + (let_op_one (If x1 x2 x3) = + If (let_op_one x1) + (let_op_one x2) + (let_op_one x3)) /\ + (let_op_one (Let xs x2) = + let xs = let_op_list xs in + let x2 = let_op_one x2 in + case dest_op x2 xs of + | SOME op => Op op xs + | NONE => Let xs x2) /\ + (let_op_one (Raise x1) = + Raise (let_op_one x1)) /\ + (let_op_one (Handle x1 x2) = + Handle (let_op_one x1) + (let_op_one x2)) /\ + (let_op_one (Op op xs) = + Op op (let_op_list xs)) /\ + (let_op_one (Tick x) = Tick (let_op_one x)) /\ + (let_op_one (Call ticks dest xs) = Call ticks dest (let_op_list xs)) /\ + (let_op_list [] = []) /\ + (let_op_list ((x:bvl$exp)::xs) = + (let_op_one x) :: let_op_list xs) +End + +Theorem let_op_one: + (∀e. let_op [e] = [let_op_one e]) ∧ + (∀es. let_op_list es = let_op es) +Proof + Induct \\ gvs [let_op_one_def,let_op_def] \\ rw [] + \\ every_case_tac \\ gvs [] + \\ Cases_on ‘es’ \\ gvs [let_op_one_def,let_op_def] +QED -val let_op_sing_def = Define ` +Definition let_op_sing_def: let_op_sing x = case let_op [x] of | (y::ys) => y - | _ => Op (Const 0) []`; + | _ => Op (Const 0) [] +End -val optimise_def = Define ` +Theorem let_op_sing_eq = let_op_sing_def |> SRULE [let_op_one]; + +Definition optimise_def: optimise split_seq cut_size (name,arity, exp) = (name,arity,bvl_handle$compile_any split_seq cut_size arity - (let_op_sing (HD (remove_ticks [exp]))))`; + (let_op_sing (HD (remove_ticks [exp])))) +End + +Theorem optimise_eq = optimise_def |> SRULE [remove_ticks_sing]; -val compile_inc_def = Define ` +Definition compile_inc_def: compile_inc limit split_seq cut_size cs prog = let (cs,prog1) = tick_compile_prog limit cs prog in - (cs, MAP (optimise split_seq cut_size) prog1)` + (cs, MAP (optimise split_seq cut_size) prog1) +End -val compile_prog_def = Define ` +Definition compile_prog_def: compile_prog limit split_seq cut_size prog = - compile_inc limit split_seq cut_size LN prog` + compile_inc limit split_seq cut_size LN prog +End val _ = export_theory(); diff --git a/compiler/backend/bvl_to_bviScript.sml b/compiler/backend/bvl_to_bviScript.sml index 3cbeedfc77..cb8af56f2c 100644 --- a/compiler/backend/bvl_to_bviScript.sml +++ b/compiler/backend/bvl_to_bviScript.sml @@ -69,6 +69,42 @@ Termination WF_REL_TAC `measure exp1_size` End +Definition global_count_sing_def: + (global_count_sing ((Var _):bvl$exp) = 0) /\ + (global_count_sing (If x y z) = + global_count_sing x + + global_count_sing y + + global_count_sing z) /\ + (global_count_sing (Handle x y) = + global_count_sing x + + global_count_sing y) /\ + (global_count_sing (Tick x) = global_count_sing x) /\ + (global_count_sing (Raise x) = global_count_sing x) /\ + (global_count_sing (Let xs x) = + global_count_sing x + global_count_list xs) /\ + (global_count_sing (Call _ _ xs) = global_count_list xs) /\ + (global_count_sing (Op op xs) = + if op = AllocGlobal then 1 + global_count_list xs + else global_count_list xs) /\ + (global_count_list [] = 0:num) /\ + (global_count_list (x::xs) = + global_count_sing x + global_count_list xs) +Termination + WF_REL_TAC ‘measure $ λx. case x of + | INL e => bvl$exp_size e + | INR es => bvl$exp1_size es’ + \\ rpt strip_tac \\ simp [bvlTheory.exp_size_def] +End + +Theorem alloc_glob_count_eq_global_count_list: + (∀e. global_count_sing e = alloc_glob_count [e]) ∧ + (∀es. global_count_list es = alloc_glob_count es) +Proof + Induct \\ gvs [global_count_sing_def,alloc_glob_count_def] + \\ rw [] \\ Cases_on ‘es’ \\ gvs [] + \\ gvs [global_count_sing_def,alloc_glob_count_def] +QED + val AllocGlobal_location_def = Define` AllocGlobal_location = data_num_stubs`; val CopyGlobals_location_def = Define` @@ -327,7 +363,75 @@ Termination \\ SRW_TAC [] [bvlTheory.exp_size_def] \\ DECIDE_TAC End -val compile_exps_ind = theorem"compile_exps_ind"; +Definition compile_exps_sing_def: + (compile_exps_sing n (Var v) = (((Var v):bvi$exp), Nil, n)) /\ + (compile_exps_sing n (If x1 x2 x3) = + let (c1,aux1,n1) = compile_exps_sing n x1 in + let (c2,aux2,n2) = compile_exps_sing n1 x2 in + let (c3,aux3,n3) = compile_exps_sing n2 x3 in + (If c1 c2 c3,aux1++aux2++aux3,n3)) /\ + (compile_exps_sing n (Let xs x2) = + if NULL xs (* i.e. a marker *) then + let (args,x0) = destLet x2 in + let (c1,aux1,n1) = compile_exps_list n args in + let (c2,aux2,n2) = compile_exps_sing n1 x0 in + let n3 = n2 + 1 in + ((Call 0 (SOME (num_stubs + nss * n2 + 1)) c1 NONE), + aux1++aux2++compile_aux(n2,LENGTH args,c2), n3) + else + let (c1,aux1,n1) = compile_exps_list n xs in + let (c2,aux2,n2) = compile_exps_sing n1 x2 in + ((Let c1 (c2)), aux1++aux2, n2)) /\ + (compile_exps_sing n (Raise x1) = + let (c1,aux1,n1) = compile_exps_sing n x1 in + (Raise c1, aux1, n1)) /\ + (compile_exps_sing n (Tick x1) = + let (c1,aux1,n1) = compile_exps_sing n x1 in + (Tick c1, aux1, n1)) /\ + (compile_exps_sing n (Op op xs) = + let (c1,aux1,n1) = compile_exps_list n xs in + (compile_op op c1,aux1,n1)) /\ + (compile_exps_sing n (Handle x1 x2) = + let (args,x0) = destLet x1 in + let (c1,aux1,n1) = compile_exps_list n args in + let (c2,aux2,n2) = compile_exps_sing n1 x0 in + let (c3,aux3,n3) = compile_exps_sing n2 x2 in + let aux4 = compile_aux(n3,LENGTH args,c2) in + let n4 = n3 + 1 in + ((Call 0 (SOME (num_stubs + nss * n3 + 1)) c1 (SOME c3)), + aux1++aux2++aux3++aux4, n4)) /\ + (compile_exps_sing n (Call ticks dest xs) = + let (c1,aux1,n1) = compile_exps_list n xs in + ((Call ticks + (dtcase dest of + | NONE => NONE + | SOME n => SOME (num_stubs + nss * n)) c1 NONE),aux1,n1)) /\ + (compile_exps_list n [] = ([],Nil,n)) /\ + (compile_exps_list n ((x:bvl$exp)::xs) = + let (c1,aux1,n1) = compile_exps_sing n x in + let (c2,aux2,n2) = compile_exps_list n1 xs in + (c1::c2, aux1 ++ aux2, n2)) +Termination + WF_REL_TAC ‘measure $ λx. case x of INL (n,e) => exp_size e + | INR (n,es) => exp1_size es’ + \\ rpt strip_tac + \\ rpt $ qpat_x_assum ‘(_,_,_) = _’ kall_tac + \\ gvs [bvlTheory.exp_size_def] + \\ gvs [oneline destLet_def, AllCaseEqs(), bvlTheory.exp_size_def] +End + +Theorem compile_exps_sing: + (∀n e. compile_exps n [e] = + (let (x,y,z) = compile_exps_sing n e in ([x],y,z))) ∧ + (∀n es. compile_exps_list n es = compile_exps n es) +Proof + ho_match_mp_tac compile_exps_sing_ind \\ rpt strip_tac + \\ gvs [compile_exps_def,compile_exps_sing_def] + \\ rpt (pairarg_tac \\ gvs []) + \\ rpt IF_CASES_TAC \\ gvs [] + \\ Cases_on ‘es’ \\ gvs [] + \\ gvs [compile_exps_def,compile_exps_sing_def,SmartAppend_Nil] +QED val compile_exps_LENGTH_lemma = Q.prove( `!n xs. (LENGTH (FST (compile_exps n xs)) = LENGTH xs)`, @@ -355,6 +459,16 @@ Definition compile_single_def: ++ aux, n1) End +Theorem compile_single_eq: + compile_single n (name,arg_count,exp) = + let (c,aux,n1) = compile_exps_sing n exp in + (List [(num_stubs + nss * name,arg_count,bvi_let$compile_exp c)] + ++ aux, n1) +Proof + gvs [compile_single_def,compile_exps_sing] + \\ rpt (pairarg_tac \\ gvs []) +QED + Definition compile_list_def: (compile_list n [] = (List [],n)) /\ (compile_list n (p::progs) = @@ -376,6 +490,9 @@ Definition compile_prog_def: (InitGlobals_location, bvl_to_bvi$stubs (num_stubs + nss * start) k ++ append code, n1) End +Theorem compile_prog_eq = + compile_prog_def |> SRULE [GSYM alloc_glob_count_eq_global_count_list]; + Datatype: config = <| inline_size_limit : num (* zero disables inlining *) ; exp_cut : num (* huge number effectively disables exp splitting *) diff --git a/compiler/backend/clos_annotateScript.sml b/compiler/backend/clos_annotateScript.sml index 3b4b55459c..66ea5d5b4b 100644 --- a/compiler/backend/clos_annotateScript.sml +++ b/compiler/backend/clos_annotateScript.sml @@ -85,6 +85,102 @@ Termination \\ simp[] End +Definition alt_free_sing_def: + (alt_free_sing (Var t v) = (Var t v, Var v)) /\ + (alt_free_sing (If t x1 x2 x3) = + let (c1,l1) = alt_free_sing x1 in + let (c2,l2) = alt_free_sing x2 in + let (c3,l3) = alt_free_sing x3 in + (If t c1 c2 c3,mk_Union l1 (mk_Union l2 l3))) /\ + (alt_free_sing (Let t xs x2) = + let m = LENGTH xs in + let (c2,l2) = alt_free_sing x2 in + if no_overlap m l2 /\ EVERY pure xs then + (Let t (REPLICATE m (const_0 t)) c2, Shift m l2) + else + let (c1,l1) = alt_free_list xs in + (Let t c1 c2,mk_Union l1 (Shift (LENGTH xs) l2))) /\ + (alt_free_sing (Raise t x1) = + let (c1,l1) = alt_free_sing x1 in + (Raise t c1,l1)) /\ + (alt_free_sing (Tick t x1) = + let (c1,l1) = alt_free_sing x1 in + (Tick t c1,l1)) /\ + (alt_free_sing (Op t op xs) = + let (c1,l1) = alt_free_list xs in + (Op t op c1,l1)) /\ + (alt_free_sing (App t loc_opt x1 xs2) = + let (c1,l1) = alt_free_sing x1 in + let (c2,l2) = alt_free_list xs2 in + (App t loc_opt c1 c2,mk_Union l1 l2)) /\ + (alt_free_sing (Fn t loc _ num_args x1) = + let (c1,l1) = alt_free_sing x1 in + let l2 = Shift num_args l1 in + (Fn t loc (SOME (vars_to_list l2)) num_args c1,l2)) /\ + (alt_free_sing (Letrec t loc _ fns x1) = + let m = LENGTH fns in + let (c2,l2) = alt_free_sing x1 in + (* if no_overlap m l2 then + ([Let t (REPLICATE m (const_0 t)) (HD c2)], Shift m l2) + else *) + let res = alt_free_letrec m fns in + let c1 = MAP FST res in + let l1 = list_mk_Union (MAP SND res) in + (Letrec t loc (SOME (vars_to_list l1)) c1 c2, + mk_Union l1 (Shift (LENGTH fns) l2))) /\ + (alt_free_sing (Handle t x1 x2) = + let (c1,l1) = alt_free_sing x1 in + let (c2,l2) = alt_free_sing x2 in + (Handle t c1 c2,mk_Union l1 (Shift 1 l2))) /\ + (alt_free_sing (Call t ticks dest xs) = + let (c1,l1) = alt_free_list xs in + (Call t ticks dest c1,l1)) ∧ + + (alt_free_list [] = ([],Empty)) ∧ + (alt_free_list ((x:closLang$exp)::xs) = + let (c1,l1) = alt_free_sing x in + let (c2,l2) = alt_free_list xs in + (c1 :: c2,mk_Union l1 l2)) ∧ + + (alt_free_letrec m [] = []) ∧ + (alt_free_letrec m ((n,x)::rest) = + let (c,l) = alt_free_sing x in + ((n,c),Shift (n + m) l) :: alt_free_letrec m rest) +Termination + WF_REL_TAC `measure $ λx. case x of INL e => exp_size e + | INR (INL es) => exp3_size es + | INR (INR (_,fns)) => exp1_size fns` +End + +Theorem alt_free_sing_eq: + (∀e. alt_free [e] = (λ(e,dbs). [e], dbs) $ alt_free_sing e) ∧ + (∀es. alt_free es = alt_free_list es) ∧ + (∀m fns. EVERY (λ(n,e). alt_free [e] = (λ(e,dbs). [e], dbs) $ alt_free_sing e) fns ∧ + alt_free_letrec m fns = + MAP (λ(n,x). let (c,l) = alt_free_sing x in + ((n,c),Shift (n + m) l)) fns) +Proof + ho_match_mp_tac alt_free_sing_ind >> reverse $ rw[alt_free_def, alt_free_sing_def] >> + rpt (pairarg_tac >> gvs[]) >> rpt (TOP_CASE_TAC >> gvs[]) + >- (Cases_on `es` >> gvs[alt_free_def, alt_free_sing_def]) >> + simp[MAP_MAP_o, combinTheory.o_DEF, ELIM_UNCURRY] >> rw[] >> + gvs [EVERY_MEM] + >- ( + ntac 2 AP_TERM_TAC >> + rw[MAP_EQ_f] >> PairCases_on `x` >> + last_x_assum drule >> rpt (pairarg_tac >> gvs[]) + ) + >- ( + rw[MAP_EQ_f] >> PairCases_on `x` >> + last_x_assum drule >> rpt (pairarg_tac >> gvs[]) + ) + >- ( + AP_THM_TAC >> ntac 2 AP_TERM_TAC >> + rw[MAP_EQ_f] >> PairCases_on `x` >> + last_x_assum drule >> rpt (pairarg_tac >> gvs[]) + ) +QED + Triviality alt_free_LENGTH_LEMMA: !xs. (case alt_free xs of (ys,s1) => (LENGTH xs = LENGTH ys)) Proof @@ -204,6 +300,88 @@ Termination \\ IMP_RES_TAC exp1_size_lemma \\ DECIDE_TAC End +Definition shift_sing_def: + (shift_sing (Var t v) m l i = + Var t (get_var m l i v)) /\ + (shift_sing (If t x1 x2 x3) m l i = + let c1 = shift_sing x1 m l i in + let c2 = shift_sing x2 m l i in + let c3 = shift_sing x3 m l i in + (If t c1 c2 c3)) /\ + (shift_sing (Let t xs x2) m l i = + let c1 = shift_list xs m l i in + let c2 = shift_sing x2 m (l + LENGTH xs) i in + (Let t c1 c2)) /\ + (shift_sing (Raise t x1) m l i = + let (c1) = shift_sing x1 m l i in + (Raise t c1)) /\ + (shift_sing (Tick t x1) m l i = + let c1 = shift_sing x1 m l i in + (Tick t c1)) /\ + (shift_sing (Op t op xs) m l i = + let c1 = shift_list xs m l i in + (Op t op c1)) /\ + (shift_sing (App t loc_opt x1 xs2) m l i = + let c1 = shift_sing x1 m l i in + let c2 = shift_list xs2 m l i in + (App t loc_opt c1 c2)) /\ + (shift_sing (Fn t loc vs_opt num_args x1) m l i = + let k = m + l in + let vs = case vs_opt of NONE => [] | SOME vs => vs in + let live = FILTER (\n. n < k) vs in + let vars = MAP (get_var m l i) live in + let c1 = shift_sing x1 k num_args (shifted_env 0 live) in + (Fn t loc (SOME vars) num_args c1)) /\ + (shift_sing (Letrec t loc vsopt fns x1) m l i = + let vs = case vsopt of NONE => [] | SOME x => x in + let k = m + l in + let live = FILTER (\n. n < k) vs in + let vars = MAP (get_var m l i) live in + let new_i = shifted_env 0 live in + let fns_len = LENGTH fns in + let cs = shift_letrec fns_len k new_i fns in + let c1 = shift_sing x1 m (l + LENGTH fns) i in + (Letrec t loc (SOME vars) cs c1)) /\ + (shift_sing (Handle t x1 x2) m l i = + let c1 = shift_sing x1 m l i in + let c2 = shift_sing x2 m (l+1) i in + (Handle t c1 c2)) /\ + (shift_sing (Call t ticks dest xs) m l i = + let c1 = shift_list xs m l i in + (Call t ticks dest c1)) ∧ + + (shift_list [] (m:num) (l:num) (i:num num_map) = []) /\ + (shift_list ((x:closLang$exp)::xs) m l i = + let c1 = shift_sing x m l i in + let c2 = shift_list xs m l i in + (c1 :: c2:closLang$exp list)) ∧ + + (shift_letrec fns_len k new_i [] = []) ∧ + (shift_letrec fns_len k new_i ((n,x)::rest) = + let c = shift_sing x k (n + fns_len) new_i in + (n,c) :: shift_letrec fns_len k new_i rest) +Termination + WF_REL_TAC `measure $ λx. case x of INL (e,_,_,_) => exp_size e + | INR (INL (es,_,_,_)) => exp3_size es + | INR (INR (_,_,_,l)) => exp1_size l` +End + +Theorem shift_sing_eq: + (∀e m l i. shift [e] m l i = [shift_sing e m l i]) ∧ + (∀es m l i. shift es m l i = shift_list es m l i) ∧ + (∀fns_len k new_i fns. + EVERY (λ(n,e). shift [e] k (n + fns_len) new_i = + [shift_sing e k (n + fns_len) new_i]) fns ∧ + shift_letrec fns_len k new_i fns = + MAP (λ(n,x). let c = shift_sing x k (n + fns_len) new_i in + (n,c)) fns) +Proof + ho_match_mp_tac shift_sing_ind >> reverse $ rw[shift_def, shift_sing_def] + >- (Cases_on `es` >> gvs[shift_def, shift_sing_def]) >> + rw[MAP_EQ_f] >> pairarg_tac >> gvs[] >> + gvs [EVERY_MEM] >> res_tac >> fs [] +QED + Theorem shift_LENGTH_LEMMA: !xs m l i. LENGTH (shift xs m l i) = LENGTH xs Proof @@ -248,4 +426,36 @@ Definition compile_inc_def: compile_inc (e,aux) = (annotate 0 e,clos_annotate$compile aux) End +(* cv versions *) + +Definition annotate_sing_def: + annotate_sing arity x = shift_sing (FST (alt_free_sing x)) 0 arity LN +End + +Definition annotate_list_def: + annotate_list arity xs = shift_list (FST (alt_free_list xs)) 0 arity LN +End + +Theorem annotate_sing_eq: + annotate arity [e] = [annotate_sing arity e] ∧ + annotate arity es = annotate_list arity es +Proof + simp[annotate_def, annotate_sing_def, annotate_list_def] >> + simp[alt_free_sing_eq, shift_sing_eq] >> + pairarg_tac >> gvs[shift_sing_def] +QED + +Theorem compile_eq: + compile prog = + MAP (λ(n,args,exp). (n,args, annotate_sing args exp)) prog +Proof + simp[compile_def, annotate_sing_eq] +QED + +Theorem compile_inc_eq: + compile_inc (e,aux) = (annotate_list 0 e, clos_annotate$compile aux) +Proof + simp[compile_inc_def, annotate_sing_eq] +QED + val _ = export_theory(); diff --git a/compiler/backend/clos_callScript.sml b/compiler/backend/clos_callScript.sml index acd6e48360..003e5789a8 100644 --- a/compiler/backend/clos_callScript.sml +++ b/compiler/backend/clos_callScript.sml @@ -61,6 +61,86 @@ val free_def = tDefine "free" ` val free_ind = theorem "free_ind"; +Definition free_sing_def: + (free_sing (Var t v) = (Var t v, Var v)) /\ + (free_sing (If t x1 x2 x3) = + let (c1,l1) = free_sing x1 in + let (c2,l2) = free_sing x2 in + let (c3,l3) = free_sing x3 in + (If t c1 c2 c3,mk_Union l1 (mk_Union l2 l3))) /\ + (free_sing (Let t xs x2) = + let (c1,l1) = free_list xs in + let (c2,l2) = free_sing x2 in + (Let t c1 c2,mk_Union l1 (Shift (LENGTH xs) l2))) /\ + (free_sing (Raise t x1) = + let (c1,l1) = free_sing x1 in + (Raise t c1,l1)) /\ + (free_sing (Tick t x1) = + let (c1,l1) = free_sing x1 in + (Tick t c1,l1)) /\ + (free_sing (Op t op xs) = + let (c1,l1) = free_list xs in + (Op t op c1,l1)) /\ + (free_sing (App t loc_opt x1 xs2) = + let (c1,l1) = free_sing x1 in + let (c2,l2) = free_list xs2 in + (App t loc_opt c1 c2,mk_Union l1 l2)) /\ + (free_sing (Fn t loc _ num_args x1) = + let (c1,l1) = free_sing x1 in + let l2 = Shift num_args l1 in + (Fn t loc (SOME (vars_to_list l2)) num_args c1,l2)) /\ + (free_sing (Letrec t loc _ fns x1) = + let m = LENGTH fns in + let res = MAP (\(n,x). let (c,l) = free_sing x in + ((n,c),Shift (n + m) l)) fns in + let c1 = MAP FST res in + let l1 = list_mk_Union (MAP SND res) in + let (c2,l2) = free_sing x1 in + (Letrec t loc (SOME (vars_to_list l1)) c1 c2, + mk_Union l1 (Shift (LENGTH fns) l2))) /\ + (free_sing (Handle t x1 x2) = + let (c1,l1) = free_sing x1 in + let (c2,l2) = free_sing x2 in + (Handle t c1 c2,mk_Union l1 (Shift 1 l2))) /\ + (free_sing (Call t ticks dest xs) = + let (c1,l1) = free_list xs in + (Call t ticks dest c1,l1)) ∧ + + (free_list [] = ([],Empty)) /\ + (free_list ((x:closLang$exp)::xs) = + let (c1,l1) = free_sing x in + let (c2,l2) = free_list xs in + (c1 :: c2,mk_Union l1 l2)) +Termination + WF_REL_TAC `measure $ λx. case x of INL e => exp_size e + | INR es => exp3_size es` +End + +Theorem free_sing_eq: + (∀e. free [e] = (λ(e,vs). ([e],vs)) $ free_sing e) ∧ + (∀es. free es = free_list es) +Proof + ho_match_mp_tac free_sing_ind >> + reverse $ rw[free_def, free_sing_def] >> + rpt (pairarg_tac >> gvs[]) + >- (Cases_on `es` >> gvs[free_def, free_sing_def]) >> + simp[MAP_MAP_o, combinTheory.o_DEF, ELIM_UNCURRY] >> rw[] + >- ( + ntac 2 AP_TERM_TAC >> + rw[MAP_EQ_f] >> PairCases_on `x` >> + last_x_assum drule >> pairarg_tac >> gvs[] + ) + >- ( + rw[MAP_EQ_f] >> PairCases_on `x` >> + last_x_assum drule >> pairarg_tac >> gvs[] + ) + >- ( + AP_THM_TAC >> ntac 2 AP_TERM_TAC >> + rw[MAP_EQ_f] >> PairCases_on `x` >> + last_x_assum drule >> pairarg_tac >> gvs[] + ) +QED + val free_LENGTH_LEMMA = Q.prove( `!xs. (case free xs of (ys,s1) => (LENGTH xs = LENGTH ys))`, recInduct free_ind \\ REPEAT STRIP_TAC @@ -108,6 +188,12 @@ QED val closed_def = Define ` closed x = isEmpty (db_to_set (SND (free [x])))` +Theorem closed_eq: + closed x = isEmpty (db_to_set (SND (free_sing x))) +Proof + simp[closed_def, free_sing_eq] >> pairarg_tac >> gvs[] +QED + val EL_MEM_LEMMA = Q.prove( `!xs i x. i < LENGTH xs /\ (x = EL i xs) ==> MEM x xs`, Induct \\ fs [] \\ REPEAT STRIP_TAC \\ Cases_on `i` \\ fs []); @@ -122,12 +208,12 @@ val code_list_def = Define ` code_list (loc+2n) xs (g1,(loc+1,n,p)::g2))` val GENLIST_Var_def = Define ` - GENLIST_Var t i n = + GENLIST_Var t (i:num) n = if n = 0 then [] else GENLIST_Var t (i+1) (n-1:num) ++ [Var t (n-1)]`; val calls_list_def = Define ` - (calls_list t i loc [] = []) /\ + (calls_list t (i:num) loc [] = []) /\ (calls_list t i loc ((n,_)::xs) = (n,Call t 0 (loc+1) (GENLIST_Var t 1 n)):: calls_list t (i+1) (loc+2n) xs)`; @@ -221,6 +307,95 @@ val calls_def = tDefine "calls" ` \\ assume_tac (SPEC_ALL exp3_size_MAP_SND) \\ DECIDE_TAC); +Definition calls_sing_def: + (calls_sing (Var t v) g = + (Var t v,g)) /\ + (calls_sing (If t x1 x2 x3) g = + let (e1,g) = calls_sing x1 g in + let (e2,g) = calls_sing x2 g in + let (e3,g) = calls_sing x3 g in + (If t e1 e2 e3,g)) /\ + (calls_sing (Let t xs x2) g = + let (e1,g) = calls_sing_list xs g in + let (e2,g) = calls_sing x2 g in + (Let t e1 e2,g)) /\ + (calls_sing (Raise t x1) g = + let (e1,g) = calls_sing x1 g in + (Raise t e1,g)) /\ + (calls_sing (Tick t x1) g = + let (e1,g) = calls_sing x1 g in + (Tick t e1,g)) /\ + (calls_sing (Handle t x1 x2) g = + let (e1,g) = calls_sing x1 g in + let (e2,g) = calls_sing x2 g in + (Handle t e1 e2,g)) /\ + (calls_sing (Call t ticks dest xs) g = + let (xs,g) = calls_sing_list xs g in + (Call t ticks dest xs,g)) /\ + (calls_sing (Op t op xs) g = + let (e1,g) = calls_sing_list xs g in + (Op t op e1,g)) /\ + (calls_sing (App t loc_opt x xs) g = + let (es,g) = calls_sing_list xs g in + let (e1,g) = calls_sing x g in + let loc = (case loc_opt of SOME loc => loc | NONE => 0) in + if IS_SOME loc_opt /\ IS_SOME (lookup loc (FST g)) then + if pure x then (Call t (LENGTH es) (loc+1) es,g) else + (Let (t§0) (SNOC e1 es) + (Call (t§1) (LENGTH es) (loc+1) (GENLIST_Var None 2 (LENGTH es))),g) + else (App t loc_opt e1 es,g)) /\ + (calls_sing (Fn t loc_opt ws num_args x1) g = + (* loc_opt ought to be SOME loc, with loc being EVEN *) + let loc = (case loc_opt of SOME loc => loc | NONE => 0) in + let new_g = insert_each loc 1 g in + let (e1,new_g) = calls_sing x1 new_g in + let new_g = (FST new_g,(loc+1,num_args,e1)::SND new_g) in + (* Closedness is checked on the transformed program because + the calls function can sometimes remove free variables. *) + if closed (Fn t loc_opt ws num_args e1) then + (Fn t loc_opt ws num_args + (Call None 0 (loc+1) (GENLIST_Var None 1 num_args)),new_g) + else + let (e1,g) = calls_sing x1 g in + (Fn t loc_opt ws num_args e1,g)) /\ + (calls_sing (Letrec t loc_opt ws fns x1) g = + let loc = (case loc_opt of SOME loc => loc | NONE => 0) in + let new_g = insert_each loc (LENGTH fns) g in + let (fns1,new_g) = calls_sing_list (MAP SND fns) new_g in + if EVERY2 (\(n,_) p. closed (Fn (strlit "") NONE NONE n p)) fns fns1 then + let new_g = code_list loc (ZIP (MAP FST fns,fns1)) new_g in + let (e1,g) = calls_sing x1 new_g in + (Letrec t loc_opt ws (calls_list None 1 loc fns) e1,g) + else + let (fns1,g) = calls_sing_list (MAP SND fns) g in + let (e1,g) = calls_sing x1 g in + (Letrec t loc_opt ws (ZIP (MAP FST fns,fns1)) e1,g)) ∧ + + (calls_sing_list [] g = ([],g)) /\ + (calls_sing_list ((x:closLang$exp)::xs) g = + let (e1,g) = calls_sing x g in + let (e2,g) = calls_sing_list xs g in + (e1 :: e2,g)) +Termination + WF_REL_TAC `measure $ λx. case x of INL (e,_) => exp_size e + | INR (es,_) => exp3_size es` + \\ REPEAT STRIP_TAC + \\ fs [GSYM NOT_LESS] + \\ IMP_RES_TAC EL_MEM_LEMMA + \\ IMP_RES_TAC exp1_size_lemma + \\ assume_tac (SPEC_ALL exp3_size_MAP_SND) + \\ DECIDE_TAC +End + +Theorem calls_sing_eq: + (∀e g. calls [e] g = (λ(e,rest). ([e], rest)) $ calls_sing e g) ∧ + (∀es g. calls es g = calls_sing_list es g) +Proof + ho_match_mp_tac calls_sing_ind >> rw[calls_def, calls_sing_def] >> + rpt (pairarg_tac >> gvs[]) >> rpt (TOP_CASE_TAC >> gvs[]) >> + Cases_on `es` >> gvs[calls_def, calls_sing_def] +QED + val compile_def = Define ` compile F x = (x,(LN,[])) /\ compile T x = let (xs,g) = calls x (LN,[]) in (xs,g)` @@ -230,6 +405,21 @@ val compile_inc_def = Define ` let (ea, d1, new_code) = calls e (d,[]) in (d1, ea, new_code)`; +Theorem compile_eq: + compile F x = (x,(LN,[])) ∧ + compile T x = let (xs,g) = calls_sing_list x (LN,[]) in (xs,g) +Proof + simp[compile_def, calls_sing_eq] +QED + +Theorem compile_inc_eq: + compile_inc d (e,xs) = + let (ea, d1, new_code) = calls_sing_list e (d,[]) in + (d1, ea, new_code) +Proof + simp[compile_inc_def, calls_sing_eq] +QED + val cond_call_compile_inc_def = Define` cond_call_compile_inc do_it = if do_it then compile_inc else CURRY I`; diff --git a/compiler/backend/clos_fvsScript.sml b/compiler/backend/clos_fvsScript.sml index fcab0dd868..1ec438c13b 100644 --- a/compiler/backend/clos_fvsScript.sml +++ b/compiler/backend/clos_fvsScript.sml @@ -48,9 +48,73 @@ val remove_fvs_def = tDefine "remove_fvs" ` \\ imp_res_tac exp1_size_lemma \\ simp []); +Definition remove_fvs_sing_def: + (remove_fvs_sing fvs (Var t v) = + if v < fvs then Var t v + else Op t El []) /\ + (remove_fvs_sing fvs (If t x1 x2 x3) = + If t (remove_fvs_sing fvs x1) + (remove_fvs_sing fvs x2) + (remove_fvs_sing fvs x3)) /\ + (remove_fvs_sing fvs (Let t xs x2) = + let xs = remove_fvs_list fvs xs in + let n = LENGTH xs in + let x2 = remove_fvs_sing (fvs+n) x2 in + Let t xs x2) /\ + (remove_fvs_sing fvs (Raise t x1) = + Raise t (remove_fvs_sing fvs x1)) /\ + (remove_fvs_sing fvs (Handle t x1 x2) = + Handle t (remove_fvs_sing fvs x1) + (remove_fvs_sing (fvs+1) x2)) /\ + (remove_fvs_sing fvs (Op t op xs) = + Op t op (remove_fvs_list fvs xs)) /\ + (remove_fvs_sing fvs (Tick t x) = Tick t (remove_fvs_sing fvs x)) /\ + (remove_fvs_sing fvs (Call t ticks dest xs) = + Call t ticks dest (remove_fvs_list fvs xs)) /\ + (remove_fvs_sing fvs (App t loc_opt x1 xs) = + App t loc_opt (remove_fvs_sing fvs x1) (remove_fvs_list fvs xs)) /\ + (remove_fvs_sing fvs (Letrec t loc_opt vs fns x1) = + let m = LENGTH fns in + let new_fvs = case vs of NONE => fvs | SOME x => LENGTH x in + let new_fns = remove_fvs_let m new_fvs fns in + Letrec t loc_opt vs new_fns (remove_fvs_sing (m + fvs) x1)) /\ + (remove_fvs_sing fvs (Fn t loc_opt vs num_args x1) = + let fvs = case vs of NONE => fvs | SOME x => LENGTH x in + Fn t loc_opt vs num_args (remove_fvs_sing (num_args+fvs) x1)) ∧ + (remove_fvs_list fvs [] = []) /\ + (remove_fvs_list fvs ((x:closLang$exp)::xs) = + remove_fvs_sing fvs x :: remove_fvs_list fvs xs) ∧ + (remove_fvs_let m fvs [] = []) /\ + (remove_fvs_let m new_fvs ((num_args,x)::xs) = + (num_args, remove_fvs_sing (num_args+m+new_fvs) x)::remove_fvs_let m new_fvs xs) +Termination + WF_REL_TAC `measure $ λx. case x of INL (_,e) => exp_size e + | INR(INL (_,es)) => exp3_size es + | INR(INR (_,_,es)) => exp1_size es` +End + +Theorem remove_fvs_sing_eq: + (∀fvs e. remove_fvs fvs [e] = [remove_fvs_sing fvs e]) ∧ + (∀fvs es. remove_fvs fvs es = remove_fvs_list fvs es) ∧ + (∀m new_fvs fns. + remove_fvs_let m new_fvs fns = + MAP (λ(num_args, x). (num_args, HD (remove_fvs (num_args+m+new_fvs) [x]))) fns + ) +Proof + ho_match_mp_tac remove_fvs_sing_ind >> + reverse $ rw[remove_fvs_def, remove_fvs_sing_def] >> + Cases_on `es` >> gvs[remove_fvs_def, remove_fvs_sing_def] +QED + val compile_def = Define` compile exps = remove_fvs 0 exps`; +Theorem compile_eq: + compile exps = remove_fvs_list 0 exps +Proof + simp[compile_def, remove_fvs_sing_eq] +QED + val compile_inc_def = Define ` compile_inc (e, xs) = (compile e, [])`; diff --git a/compiler/backend/clos_letopScript.sml b/compiler/backend/clos_letopScript.sml index a43f8eaeca..f18450d882 100644 --- a/compiler/backend/clos_letopScript.sml +++ b/compiler/backend/clos_letopScript.sml @@ -65,6 +65,59 @@ Proof \\ rw [] \\ CASE_TAC \\ simp [] QED +Definition let_op_sing_def: + (let_op_sing (Var t v) = (Var t v)) /\ + (let_op_sing (If t x1 x2 x3) = + (If t ( (let_op_sing (x1))) + ( (let_op_sing (x2))) + ( (let_op_sing (x3))))) /\ + (let_op_sing (Let t xs x2) = + let xs = let_op_list xs in + let x2 = (let_op_sing (x2)) in + case dest_op x2 xs of + | SOME op => (Op t op xs) + | NONE => (Let t xs x2)) /\ + (let_op_sing (Raise t x1) = + (Raise t ( (let_op_sing (x1))))) /\ + (let_op_sing (Handle t x1 x2) = + (Handle t ( (let_op_sing (x1))) + ( (let_op_sing (x2))))) /\ + (let_op_sing (Op t op xs) = + (Op t op (let_op_list xs))) /\ + (let_op_sing (Tick t x) = (Tick t ( (let_op_sing (x))))) /\ + (let_op_sing (Call t ticks dest xs) = (Call t ticks dest (let_op_list xs))) /\ + (let_op_sing (App t loc_opt x1 xs) = (App t loc_opt ( (let_op_sing (x1))) (let_op_list xs))) /\ + (let_op_sing (Letrec t loc_opt vs fns x1) = + let new_fns = let_op_let fns in + (Letrec t loc_opt vs new_fns ( (let_op_sing (x1))))) /\ + (let_op_sing (Fn t loc_opt vs num_args x1) = (Fn t loc_opt vs num_args ( (let_op_sing (x1))))) ∧ + (let_op_list [] = []) /\ + (let_op_list ((x:closLang$exp)::xs) = + let_op_sing x :: let_op_list xs) ∧ + (let_op_let [] = []) /\ + (let_op_let ((num_args,x:closLang$exp)::xs) = + (num_args, let_op_sing x) :: let_op_let xs) +Termination + wf_rel_tac ‘measure $ λx. case x of + | INL x => exp_size x + | INR(INL x) => exp3_size x + | INR(INR x) => exp1_size x’ +End + +Theorem let_op_sing_eq: + (∀e. let_op [e] = [let_op_sing e]) ∧ + (∀es. let_op es = let_op_list es) ∧ + (∀fns. + let_op_let fns = + MAP (λ(num_args, x). (num_args, HD (let_op [x]))) fns + ) +Proof + ho_match_mp_tac let_op_sing_ind >> + reverse $ rw[let_op_def, let_op_sing_def] + >- (Cases_on `es` >> gvs[let_op_def, let_op_sing_def]) >> + rpt(PURE_TOP_CASE_TAC >> gvs[]) +QED + val compile_inc_def = Define ` compile_inc (e, xs) = (let_op e, [])`; diff --git a/compiler/backend/clos_mtiScript.sml b/compiler/backend/clos_mtiScript.sml index 67b40d81fd..47011c8268 100644 --- a/compiler/backend/clos_mtiScript.sml +++ b/compiler/backend/clos_mtiScript.sml @@ -18,15 +18,17 @@ val collect_args_def = Define ` val collect_args_ind = theorem "collect_args_ind"; -val collect_args_size = Q.prove ( - `!max_app num_args e num_args' e'. +Theorem collect_args_size: + !max_app num_args e num_args' e'. (num_args', e') = collect_args max_app num_args e ⇒ - num_args' + exp_size e' ≤ num_args + exp_size e`, + num_args' + exp_size e' ≤ num_args + exp_size e +Proof ho_match_mp_tac collect_args_ind >> srw_tac[][collect_args_def, exp_size_def] >> srw_tac[][exp_size_def] >> res_tac >> - decide_tac); + decide_tac +QED val collect_args_more = Q.prove ( `!max_app num_args e num_args' e'. @@ -66,17 +68,19 @@ val exp3_size_append = Q.prove ( Induct_on `es1` >> simp [exp_size_def]); -val collect_apps_size = Q.prove ( - `!max_app args e args' e'. +Theorem collect_apps_size: + !max_app args e args' e'. (args', e') = collect_apps max_app args e ⇒ - exp3_size args' + exp_size e' ≤ exp3_size args + exp_size e`, + exp3_size args' + exp_size e' ≤ exp3_size args + exp_size e +Proof ho_match_mp_tac collect_apps_ind >> simp [collect_apps_def, exp_size_def, basicSizeTheory.option_size_def] >> srw_tac[][] >> simp [exp_size_def, basicSizeTheory.option_size_def] >> res_tac >> full_simp_tac(srw_ss())[exp_size_def, exp3_size_append] >> - decide_tac); + decide_tac +QED val collect_apps_more = Q.prove ( `!max_app args e args' e'. diff --git a/compiler/backend/clos_opScript.sml b/compiler/backend/clos_opScript.sml index 778b2d237e..27bb3e19ee 100644 --- a/compiler/backend/clos_opScript.sml +++ b/compiler/backend/clos_opScript.sml @@ -259,7 +259,7 @@ Termination \\ rw [] \\ res_tac \\ fs [] End -Triviality cons_measure_lemma: +Theorem cons_measure_lemma: (dest_Op_Cons x = SOME (t,xs) ⇒ cons_measure x = SUM (MAP cons_measure xs) + LENGTH xs + 1) ∧ (dest_Op_Cons x = NONE ⇒ cons_measure x = 0) ∧ (cons_measure (Op None (ElemAt i) [y]) = 0) diff --git a/compiler/backend/clos_ticksScript.sml b/compiler/backend/clos_ticksScript.sml index f5c9466e1b..afa886b5e9 100644 --- a/compiler/backend/clos_ticksScript.sml +++ b/compiler/backend/clos_ticksScript.sml @@ -45,6 +45,44 @@ val remove_ticks_def = tDefine "remove_ticks" ` \\ imp_res_tac exp1_size_lemma \\ simp []); +Definition remove_ticks_sing_def: + (remove_ticks_sing (Var t v) = (Var t v)) /\ + (remove_ticks_sing (If t x1 x2 x3) = + (If t ( (remove_ticks_sing (x1))) + ( (remove_ticks_sing (x2))) + ( (remove_ticks_sing (x3))))) /\ + (remove_ticks_sing (Let t xs x2) = + (Let t (remove_ticks_list xs) ( (remove_ticks_sing (x2))))) /\ + (remove_ticks_sing (Raise t x1) = + (Raise t ( (remove_ticks_sing (x1))))) /\ + (remove_ticks_sing (Handle t x1 x2) = + (Handle t ( (remove_ticks_sing (x1))) + ( (remove_ticks_sing (x2))))) /\ + (remove_ticks_sing (Op t op xs) = + (Op t op (remove_ticks_list xs))) /\ + (remove_ticks_sing (Tick t x1) = remove_ticks_sing (x1)) /\ + (remove_ticks_sing (Call t ticks dest xs) = + (Call t 0 dest (remove_ticks_list xs))) /\ + (remove_ticks_sing (App t loc_opt x1 xs) = + (App t loc_opt ( (remove_ticks_sing (x1))) (remove_ticks_list xs))) /\ + (remove_ticks_sing (Letrec t loc_opt vs fns x1) = + let new_fns = remove_ticks_let fns in + (Letrec t loc_opt vs new_fns ( (remove_ticks_sing (x1))))) /\ + (remove_ticks_sing (Fn t loc_opt vs num_args x1) = + (Fn t loc_opt vs num_args ( (remove_ticks_sing (x1))))) ∧ + (remove_ticks_list [] = []) /\ + (remove_ticks_list ((x:closLang$exp)::xs) = + remove_ticks_sing x :: remove_ticks_list xs) ∧ + (remove_ticks_let [] = []) /\ + (remove_ticks_let ((num_args, x)::xs) = + (num_args, remove_ticks_sing x) :: remove_ticks_let xs) +Termination + wf_rel_tac ‘measure $ λx. case x of + | INL x => exp_size x + | INR(INL x) => exp3_size x + | INR(INR x) => exp1_size x’ +End + val remove_ticks_ind = theorem "remove_ticks_ind"; Theorem LENGTH_remove_ticks: @@ -53,6 +91,19 @@ Proof recInduct remove_ticks_ind \\ fs [remove_ticks_def] QED +Theorem remove_ticks_sing_eq: + (∀e. remove_ticks [e] = [remove_ticks_sing e]) ∧ + (∀es. remove_ticks es = remove_ticks_list es) ∧ + (∀fns. + remove_ticks_let fns = + MAP (λ(num_args, x). (num_args, HD (remove_ticks [x]))) fns + ) +Proof + ho_match_mp_tac remove_ticks_sing_ind >> + reverse $ rw[remove_ticks_def, remove_ticks_sing_def] >> + Cases_on `es` >> gvs[remove_ticks_def, remove_ticks_sing_def] +QED + val compile_inc_def = Define ` compile_inc (e, xs) = (remove_ticks e, [])`; diff --git a/compiler/backend/clos_to_bvlScript.sml b/compiler/backend/clos_to_bvlScript.sml index 7bf8fdd5ee..f70a593fb8 100644 --- a/compiler/backend/clos_to_bvlScript.sml +++ b/compiler/backend/clos_to_bvlScript.sml @@ -466,12 +466,119 @@ val compile_exps_def = tDefine "compile_exps" ` val compile_exps_ind = theorem"compile_exps_ind"; +Definition compile_exp_sing_def: + (compile_exp_sing max_app (Var t v) aux = ((Var v):bvl$exp, aux)) /\ + (compile_exp_sing max_app (If t x1 x2 x3) aux = + let (c1,aux1) = compile_exp_sing max_app x1 aux in + let (c2,aux2) = compile_exp_sing max_app x2 aux1 in + let (c3,aux3) = compile_exp_sing max_app x3 aux2 in + (If c1 c2 c3,aux3)) /\ + (compile_exp_sing max_app (Let t xs x2) aux = + let (c1,aux1) = compile_exp_list max_app xs aux in + let (c2,aux2) = compile_exp_sing max_app x2 aux1 in + (Let c1 c2, aux2)) /\ + (compile_exp_sing max_app (Raise t x1) aux = + let (c1,aux1) = compile_exp_sing max_app x1 aux in + (Raise c1, aux1)) /\ + (compile_exp_sing max_app (Tick t x1) aux = + let (c1,aux1) = compile_exp_sing max_app x1 aux in + (Tick c1, aux1)) /\ + (compile_exp_sing max_app (Op t op xs) aux = + let (c1,aux1) = compile_exp_list max_app xs aux in + (if op = Install then Call 0 NONE [Op Install c1] + else Op (compile_op op) c1 + ,aux1)) /\ + (compile_exp_sing max_app (App t loc_opt x1 xs2) aux = + let (c1,aux1) = compile_exp_sing max_app x1 aux in + let (c2,aux2) = compile_exp_list max_app xs2 aux1 in + ((dtcase loc_opt of + | NONE => + Let (c2++[c1]) (mk_cl_call (Var (LENGTH c2)) (GENLIST Var (LENGTH c2))) + | SOME loc => + (Call (LENGTH c2 - 1) (SOME (loc + (num_stubs max_app))) (c2 ++ [c1]))), + aux2)) /\ + (compile_exp_sing max_app (Fn t loc_opt vs_opt num_args x1) aux = + let loc = dtcase loc_opt of NONE => 0 | SOME n => n in + let vs = dtcase vs_opt of NONE => [] | SOME vs => vs in + let (c1,aux1) = compile_exp_sing max_app x1 aux in + let c2 = + Let (GENLIST Var num_args ++ free_let (Var num_args) (LENGTH vs)) c1 + in + (Op (Cons closure_tag) + (REVERSE (mk_label (loc + num_stubs max_app) :: mk_const (num_args - 1) :: MAP Var vs)), + (loc + (num_stubs max_app),num_args+1,c2) :: aux1)) /\ + (compile_exp_sing max_app (Letrec t loc_opt vsopt fns x1) aux = + let loc = dtcase loc_opt of NONE => 0 | SOME n => n in + let vs = dtcase vsopt of NONE => [] | SOME x => x in + dtcase fns of + | [] => compile_exp_sing max_app x1 aux + | [(num_args, exp)] => + let (c1,aux1) = compile_exp_sing max_app exp aux in + let c3 = Let (GENLIST Var num_args ++ [Var num_args] ++ free_let (Var num_args) (LENGTH vs)) c1 in + let (c2,aux2) = compile_exp_sing max_app x1 ((loc + (num_stubs max_app),num_args+1,c3)::aux1) in + let c4 = + Op (Cons closure_tag) + (REVERSE (mk_label (loc + (num_stubs max_app)) :: mk_const (num_args - 1) :: MAP Var vs)) + in + (bvl$Let [c4] c2, aux2) + | _ => + let fns_l = LENGTH fns in + let l = fns_l + LENGTH vs in + let (cs,aux1) = compile_exp_list max_app (MAP SND fns) aux in + let cs1 = MAP2 (code_for_recc_case l) (MAP FST fns) cs in + let (n2,aux2) = build_aux (loc + (num_stubs max_app)) cs1 aux1 in + let (c3,aux3) = compile_exp_sing max_app x1 aux2 in + let c4 = build_recc_lets (MAP FST fns) vs (loc + (num_stubs max_app)) fns_l c3 in + (c4,aux3)) /\ + (compile_exp_sing max_app (Handle t x1 x2) aux = + let (c1,aux1) = compile_exp_sing max_app x1 aux in + let (c2,aux2) = compile_exp_sing max_app x2 aux1 in + (Handle c1 c2, aux2)) /\ + (compile_exp_sing max_app (Call t ticks dest xs) aux = + let (c1,aux1) = compile_exp_list max_app xs aux in + (Call ticks (SOME (dest + (num_stubs max_app))) c1,aux1)) ∧ + + (compile_exp_list max_app [] aux = ([],aux)) /\ + (compile_exp_list max_app ((x:closLang$exp)::xs) aux = + let (c1,aux1) = compile_exp_sing max_app x aux in + let (c2,aux2) = compile_exp_list max_app xs aux1 in + (c1 :: c2, aux2)) +Termination + WF_REL_TAC `measure $ λx. case x of INL (_,e,_) => exp_size e + | INR (_,es,_) => exp3_size es` >> + srw_tac [ARITH_ss] [closLangTheory.exp_size_def] >> + `!l. closLang$exp3_size (MAP SND l) <= exp1_size l` by ( + Induct_on `l` >> rw [closLangTheory.exp_size_def] >> + PairCases_on `h` >> full_simp_tac (srw_ss()++ARITH_ss) [closLangTheory.exp_size_def]) >> + pop_assum (qspec_then `v7` assume_tac) >> decide_tac +End + +Theorem compile_exp_sing_eq: + (∀n e l. compile_exps n [e] l = (λ(a,b). ([a], b)) (compile_exp_sing n e l)) ∧ + (∀n es l. compile_exps n es l = compile_exp_list n es l) +Proof + ho_match_mp_tac compile_exp_sing_ind >> + reverse $ rw[compile_exps_def, compile_exp_sing_def] >> + rpt (pairarg_tac >> gvs[]) + >- (Cases_on `es` >> simp[compile_exps_def] >> gvs[compile_exp_sing_def]) >> + `compile_exps n (MAP SND fns) l = compile_exp_list n (MAP SND fns) l` by ( + namedCases_on `fns` ["", "fnsh fnst"] >> + gvs[compile_exps_def, compile_exp_sing_def] >> + rpt (pairarg_tac >> gvs[]) >> + PairCases_on `fnsh` >> gvs[] >> Cases_on `fnst` >> gvs[] >> + gvs[compile_exps_def, compile_exp_sing_def]) >> + TOP_CASE_TAC >> gvs[] >> TOP_CASE_TAC >> gvs[]>> + PairCases_on `h` >> gvs[] >> rpt (pairarg_tac >> gvs[]) +QED + val compile_prog_def = Define` compile_prog max_app prog = let (new_exps, aux) = compile_exps max_app (MAP (SND o SND) prog) [] in MAP2 (λ(loc,args,_) exp. (loc + num_stubs max_app, args, exp)) prog new_exps ++ aux`; +Theorem compile_prog_eq = compile_prog_def |> SRULE [compile_exp_sing_eq]; + val pair_lem1 = Q.prove ( `!f x. (\(a,b). f a b) x = f (FST x) (SND x)`, rw [] >> @@ -666,6 +773,49 @@ Termination \\ fs [FORALL_PROD,closLangTheory.exp_size_def] End +Definition get_src_names_sing_def: + get_src_names_sing (closLang$If _ x y z) l = + get_src_names_sing x (get_src_names_sing y (get_src_names_sing z l)) ∧ + get_src_names_sing (closLang$Var _ _) l = l ∧ + get_src_names_sing (closLang$Let _ xs x) l = + get_src_names_list xs (get_src_names_sing x l) ∧ + get_src_names_sing (Raise _ x) l = get_src_names_sing x l ∧ + get_src_names_sing (Handle _ x y) l = get_src_names_sing x (get_src_names_sing y l) ∧ + get_src_names_sing (Tick _ x) l = get_src_names_sing x l ∧ + get_src_names_sing (Call _ _ _ xs) l = get_src_names_list xs l ∧ + get_src_names_sing (Op _ _ xs) l = get_src_names_list xs l ∧ + get_src_names_sing (App _ _ x xs) l = + get_src_names_list xs (get_src_names_sing x l) ∧ + get_src_names_sing (Fn name loc_opt _ _ x) l = + (let l1 = get_src_names_sing x l in + dtcase loc_opt of NONE => l1 + | SOME n => add_src_names n [name] l1) ∧ + get_src_names_sing (Letrec names loc_opt _ funs x) l = + (let l0 = get_src_names_list (MAP SND funs) l in + let l1 = get_src_names_sing x l0 in + dtcase loc_opt of NONE => l1 + | SOME n => add_src_names n names l1) ∧ + + get_src_names_list [] l = l ∧ + get_src_names_list (x::xs) l = get_src_names_list xs (get_src_names_sing x l) +Termination + WF_REL_TAC ‘measure $ λx. case x of INL (e,_) => closLang$exp_size e + | INR (es,_) => closLang$exp3_size es’ + \\ rw [closLangTheory.exp_size_def] + \\ qsuff_tac ‘exp3_size (MAP SND funs) <= exp1_size funs’ \\ fs [] + \\ Induct_on ‘funs’ \\ fs [closLangTheory.exp_size_def] + \\ fs [FORALL_PROD,closLangTheory.exp_size_def] +End + +Theorem get_src_names_sing_eq: + (∀e l. get_src_names [e] l = get_src_names_sing e l) ∧ + (∀es l. get_src_names es l = get_src_names_list es l) +Proof + ho_match_mp_tac get_src_names_sing_ind >> + rw[get_src_names_def, get_src_names_sing_def] >> + Cases_on `es` >> rw[get_src_names_def, get_src_names_sing_def] +QED + Definition make_name_alist_def: make_name_alist nums prog nstubs dec_start (dec_length:num) = let src_names = get_src_names (MAP (SND o SND) prog) LN in @@ -681,6 +831,9 @@ Definition make_name_alist_def: : mlstring$mlstring num_map End +Theorem make_name_alist_eq = + make_name_alist_def |> SRULE [get_src_names_sing_eq]; + val compile_def = Define ` compile c0 es = let (c, prog) = compile_common c0 es in diff --git a/compiler/backend/cv_compute/backend_asmScript.sml b/compiler/backend/cv_compute/backend_asmScript.sml index 073dbdc901..57956449fa 100644 --- a/compiler/backend/cv_compute/backend_asmScript.sml +++ b/compiler/backend/cv_compute/backend_asmScript.sml @@ -212,11 +212,16 @@ Definition lab_to_target_def: End Definition attach_bitmaps_def: - attach_bitmaps names (c:inc_config) bm (SOME (bytes,c')) = - SOME (bytes,bm, encode_backend_config $ c with - <| inc_lab_conf := c'; - inc_symbols := MAP (λ(n,p,l). (lookup_any n names «NOTFOUND»,p,l)) - c'.inc_sec_pos_len |>) ∧ + attach_bitmaps names (c:inc_config) data (SOME (code_bytes,c')) = + (let ffi_names = ffinames_to_string_list (the [] c'.inc_ffi_names) in + let syms = MAP (λ(n,p,l). (lookup_any n names «NOTFOUND»,p,l)) + c'.inc_sec_pos_len + in + SOME (code_bytes, LENGTH code_bytes, + data, LENGTH data, + ffi_names, LENGTH c'.inc_shmem_extra, + syms, encode_backend_config $ c with + <| inc_lab_conf := c'; inc_symbols := syms |>)) ∧ attach_bitmaps names c bm NONE = NONE End @@ -306,9 +311,16 @@ End *----------------------------------------------------------------*) Theorem from_lab_thm[local]: - from_lab asm_conf c names p bm = SOME (bytes,bm1,conf_str) ⇒ + from_lab asm_conf c names p bm = + SOME (bytes,bytes_len,bm1,bm1_len,ffi_names,shmem_len,syms,conf_str) ⇒ ∃c1. - backend$from_lab (inc_config_to_config asm_conf c) names p bm = SOME (bytes,bm1,c1) ∧ + backend$from_lab (inc_config_to_config asm_conf c) names p bm = + SOME (bytes,bm1,c1) ∧ + ffi_names = ffinames_to_string_list (the [] c1.lab_conf.ffi_names) ∧ + syms = c1.symbols ∧ + LENGTH bytes = bytes_len ∧ + LENGTH bm1 = bm1_len ∧ + LENGTH c1.lab_conf.shmem_extra = shmem_len ∧ conf_str = encode_backend_config (config_to_inc_config c1) Proof gvs [from_lab_def,backendTheory.from_lab_def] @@ -328,9 +340,15 @@ Proof QED Theorem from_stack_thm[local]: - from_stack asm_conf c names p bm = SOME (bytes,bm1,conf_str) ⇒ + from_stack asm_conf c names p bm = + SOME (bytes,bytes_len,bm1,bm1_len,ffi_names,shmem_len,syms,conf_str) ⇒ ∃c1. backend$from_stack (inc_config_to_config asm_conf c) names p bm = SOME (bytes,bm1,c1) ∧ + ffi_names = ffinames_to_string_list (the [] c1.lab_conf.ffi_names) ∧ + syms = c1.symbols ∧ + LENGTH bytes = bytes_len ∧ + LENGTH bm1 = bm1_len ∧ + LENGTH c1.lab_conf.shmem_extra = shmem_len ∧ conf_str = encode_backend_config (config_to_inc_config c1) Proof gvs [from_stack_def,backendTheory.from_stack_def] \\ rw [] @@ -360,9 +378,15 @@ Proof QED Theorem from_word_0_thm[local]: - from_word_0 asm_conf (c,p,names) = SOME (bytes,bm,conf_str) ⇒ + from_word_0 asm_conf (c,p,names) = + SOME (bytes,bytes_len,bm1,bm1_len,ffi_names,shmem_len,syms,conf_str) ⇒ ∃c1. - backend$from_word_0 (inc_config_to_config asm_conf c) names p = SOME (bytes,bm,c1) ∧ + backend$from_word_0 (inc_config_to_config asm_conf c) names p = SOME (bytes,bm1,c1) ∧ + ffi_names = ffinames_to_string_list (the [] c1.lab_conf.ffi_names) ∧ + syms = c1.symbols ∧ + LENGTH bytes = bytes_len ∧ + LENGTH bm1 = bm1_len ∧ + LENGTH c1.lab_conf.shmem_extra = shmem_len ∧ conf_str = encode_backend_config (config_to_inc_config c1) Proof gvs [from_word_0_def,from_word_def,AllCaseEqs()] \\ strip_tac \\ gvs [] @@ -465,9 +489,15 @@ QED Theorem compile_cake_thm: ∀asm_conf:'a asm_config. - compile_cake asm_conf c p = SOME (bytes,bm,conf_str) ⇒ + compile_cake asm_conf c p = + SOME (bytes,bytes_len,bm,bm_len,ffi_names,shmem_len,syms,conf_str) ⇒ ∃c1. backend$compile (inc_config_to_config asm_conf c) p = SOME (bytes,bm,c1) ∧ + ffi_names = ffinames_to_string_list (the [] c1.lab_conf.ffi_names) ∧ + syms = c1.symbols ∧ + LENGTH bytes = bytes_len ∧ + LENGTH bm = bm_len ∧ + LENGTH c1.lab_conf.shmem_extra = shmem_len ∧ conf_str = encode_backend_config (config_to_inc_config c1) Proof rw [compile_cake_def] @@ -482,4 +512,10 @@ Proof \\ rw [] \\ gvs [] QED +Theorem exists_oracle: + P x ⇒ ∃oracle. P oracle +Proof + metis_tac [] +QED + val _ = export_theory(); diff --git a/compiler/backend/proofs/backendProofScript.sml b/compiler/backend/proofs/backendProofScript.sml index 9ef8b5cded..34835b02ce 100644 --- a/compiler/backend/proofs/backendProofScript.sml +++ b/compiler/backend/proofs/backendProofScript.sml @@ -156,6 +156,14 @@ Proof rw[mc_init_ok_def] QED +Theorem set_oracle_simp[simp]: + backend_config_ok (set_oracle c x) = backend_config_ok c ∧ + mc_init_ok (backend$set_oracle c x) = mc_init_ok c ∧ + (set_oracle c x).stack_conf = c.stack_conf +Proof + gvs [backendTheory.set_oracle_def,mc_init_ok_def,FUN_EQ_THM] +QED + Theorem mc_init_ok_call_empty_ffi[simp]: mc_init_ok (cc with data_conf updated_by (λc. c with call_empty_ffi updated_by x)) = @@ -2698,7 +2706,7 @@ Proof \\ gs [] QED -Triviality compile_asm_config_eq: +Theorem compile_asm_config_eq: compile (c : 'a config) prog = SOME (b,bm,c') ==> c'.lab_conf.asm_conf = c.lab_conf.asm_conf Proof diff --git a/compiler/bootstrap/compilation/ag32/32/Holmakefile b/compiler/bootstrap/compilation/ag32/32/Holmakefile index 7c00a32487..7061d63773 100644 --- a/compiler/bootstrap/compilation/ag32/32/Holmakefile +++ b/compiler/bootstrap/compilation/ag32/32/Holmakefile @@ -1,6 +1,6 @@ ARCH = ag32 WORD_SIZE = 32 -INCLUDES = $(CAKEMLDIR)/semantics $(CAKEMLDIR)/basis ../../../..\ +INCLUDES = $(CAKEMLDIR)/semantics $(CAKEMLDIR)/basis $(CAKEMLDIR)/cv_translator ../../../..\ ../../../../encoders/asm ../../../../encoders/$(ARCH)\ ../../../../backend/$(ARCH) ../../../translation @@ -18,3 +18,7 @@ cake-$(ARCH)-$(WORD_SIZE).tar.gz: cake.S basis_ffi.c Makefile tar -chzf $@ --transform='s|^|cake-$(ARCH)-$(WORD_SIZE)/|' cake.S basis_ffi.c Makefile EXTRA_CLEANS = cake.S cake-$(ARCH)-$(WORD_SIZE).tar.gz + +ifdef POLY +HOLHEAP = $(CAKEMLDIR)/cv_translator/cake_compile_heap +endif diff --git a/compiler/bootstrap/compilation/ag32/32/README.md b/compiler/bootstrap/compilation/ag32/32/README.md index 829576cf12..b9fffc0dc8 100644 --- a/compiler/bootstrap/compilation/ag32/32/README.md +++ b/compiler/bootstrap/compilation/ag32/32/README.md @@ -13,11 +13,3 @@ library, as a thin wrapper around the relevant system calls. [proofs](proofs): This directory contains the end-to-end correctness theorem for the 32-bit version of the CakeML compiler. - -[to_dataBootstrapScript.sml](to_dataBootstrapScript.sml): -Evaluate the 32-bit version of the compiler down to a DataLang -program. - -[to_lab_ag32BootstrapScript.sml](to_lab_ag32BootstrapScript.sml): -Evaluate the 32-bit version of the compiler down to a LabLang -program (an assembly program). diff --git a/compiler/bootstrap/compilation/ag32/32/ag32BootstrapScript.sml b/compiler/bootstrap/compilation/ag32/32/ag32BootstrapScript.sml index 1dfa75ba0d..a1e5c2f89b 100644 --- a/compiler/bootstrap/compilation/ag32/32/ag32BootstrapScript.sml +++ b/compiler/bootstrap/compilation/ag32/32/ag32BootstrapScript.sml @@ -2,17 +2,11 @@ Evaluate the final part of the 32-bit version of the compiler into machine code for ag32, i.e. the Silver ISA. *) -open preamble to_lab_ag32BootstrapTheory compilationLib +open preamble compiler32ProgTheory eval_cake_compile_ag32Lib; val _ = new_theory "ag32Bootstrap"; -val () = ml_translatorLib.reset_translation(); - -val bootstrap_thm = - compilationLib.cbv_to_bytes_ag32 - stack_to_lab_thm lab_prog_def - "code" "data" "config" "cake.S"; - -val cake_compiled = save_thm("cake_compiled", bootstrap_thm); +Theorem compiler32_compiled = + eval_cake_compile_ag32 "" compiler32_prog_def "cake.S"; val _ = export_theory(); diff --git a/compiler/bootstrap/compilation/ag32/32/proofs/ag32BootstrapProofScript.sml b/compiler/bootstrap/compilation/ag32/32/proofs/ag32BootstrapProofScript.sml index d30e17bd63..709fe3bab0 100644 --- a/compiler/bootstrap/compilation/ag32/32/proofs/ag32BootstrapProofScript.sml +++ b/compiler/bootstrap/compilation/ag32/32/proofs/ag32BootstrapProofScript.sml @@ -9,17 +9,19 @@ open preamble val _ = new_theory"ag32BootstrapProof"; -val with_clos_conf_simp = prove( - ``(mc_init_ok (ag32_backend_config with <| clos_conf := z ; bvl_conf updated_by +Triviality with_clos_conf_simp: + (mc_init_ok (ag32_backend_config with <| clos_conf := z ; bvl_conf updated_by (λc. c with <|inline_size_limit := t1; exp_cut := t2|>) |>) = mc_init_ok ag32_backend_config) /\ (x.max_app <> 0 /\ (case x.known_conf of NONE => T | SOME k => k.val_approx_spt = LN) ==> (backend_config_ok (ag32_backend_config with clos_conf := x) = - backend_config_ok ag32_backend_config))``, + backend_config_ok ag32_backend_config)) +Proof fs [mc_init_ok_def,FUN_EQ_THM,backend_config_ok_def] - \\ rw [] \\ eq_tac \\ rw [] \\ EVAL_TAC); + \\ rw [] \\ eq_tac \\ rw [] \\ EVAL_TAC +QED -Overload cake_config = “ag32Bootstrap$config”; +Overload cake_config = “ag32Bootstrap$info”; Definition compiler_instance_def: compiler_instance = @@ -43,7 +45,10 @@ QED Theorem cake_config_lab_conf_asm_conf: cake_config.lab_conf.asm_conf = ag32_config Proof - once_rewrite_tac [ag32BootstrapTheory.config_def] \\ EVAL_TAC + assume_tac $ cj 1 compiler32_compiled + \\ drule compile_asm_config_eq + \\ gvs [backendTheory.set_oracle_def] + \\ strip_tac \\ EVAL_TAC QED val cake_io_events_def = new_specification("cake_io_events_def",["cake_io_events"], @@ -58,30 +63,38 @@ val cake_io_events_def = new_specification("cake_io_events_def",["cake_io_events val (cake_sem,cake_output) = cake_io_events_def |> SPEC_ALL |> UNDISCH |> CONJ_PAIR val (cake_not_fail,cake_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail cake_sem |> CONJ_PAIR -val ffi_names = - ``config.lab_conf.ffi_names`` - |> (REWRITE_CONV[ag32BootstrapTheory.config_def] THENC EVAL); - -val LENGTH_code = - ``LENGTH code`` - |> (REWRITE_CONV[ag32BootstrapTheory.code_def] THENC listLib.LENGTH_CONV); +Theorem extcalls_ffi_names: + extcalls cake_config.lab_conf.ffi_names = ffis +Proof + rewrite_tac [compiler32_compiled] + \\ qspec_tac (‘cake_config.lab_conf.ffi_names’,‘xs’) \\ Cases + \\ gvs [extcalls_def,backendTheory.ffinames_to_string_list_def, + libTheory.the_def] + \\ Induct_on ‘x’ \\ gvs [] + \\ gvs [extcalls_def,backendTheory.ffinames_to_string_list_def, + libTheory.the_def] + \\ Cases + \\ gvs [extcalls_def,backendTheory.ffinames_to_string_list_def, + libTheory.the_def] +QED -val LENGTH_data = - ``LENGTH data`` - |> (REWRITE_CONV[ag32BootstrapTheory.data_def] THENC listLib.LENGTH_CONV); +val ffis = ffis_def |> CONV_RULE (RAND_CONV EVAL); +val ffi_names = extcalls_ffi_names |> SRULE [ffis] -val shmem = - ``config.lab_conf.shmem_extra`` - |> (REWRITE_CONV[ag32BootstrapTheory.config_def] THENC EVAL); +val LENGTH_code = “LENGTH code” |> SCONV [compiler32_compiled]; +val LENGTH_data = “LENGTH data” |> SCONV [compiler32_compiled]; +val shmem = “info.lab_conf.shmem_extra” |> SCONV [compiler32_compiled]; Overload cake_machine_config = - ``ag32_machine_config (extcalls config.lab_conf.ffi_names) (LENGTH code) (LENGTH data)`` + “ag32_machine_config (extcalls info.lab_conf.ffi_names) (LENGTH code) (LENGTH data)” Theorem target_state_rel_cake_start_asm_state: SUM (MAP strlen cl) + LENGTH cl ≤ cline_size ∧ LENGTH inp ≤ stdin_size ∧ - is_ag32_init_state (init_memory code data (extcalls config.lab_conf.ffi_names) (cl,inp)) ms ⇒ - ∃n. target_state_rel ag32_target (init_asm_state code data (extcalls config.lab_conf.ffi_names) (cl,inp)) (FUNPOW Next n ms) ∧ + is_ag32_init_state (init_memory code data (extcalls info.lab_conf.ffi_names) (cl,inp)) ms ⇒ + ∃n. target_state_rel ag32_target + (init_asm_state code data (extcalls info.lab_conf.ffi_names) (cl,inp)) + (FUNPOW Next n ms) ∧ ((FUNPOW Next n ms).io_events = ms.io_events) ∧ (∀x. x ∉ (ag32_startup_addresses) ⇒ ((FUNPOW Next n ms).MEM x = ms.MEM x)) @@ -90,7 +103,7 @@ Proof \\ drule (GEN_ALL init_asm_state_RTC_asm_step) \\ disch_then drule \\ simp_tac std_ss [] - \\ disch_then(qspecl_then[`code`,`data`,`extcalls config.lab_conf.ffi_names`]mp_tac) + \\ disch_then(qspecl_then[`code`,`data`,`extcalls info.lab_conf.ffi_names`]mp_tac) \\ impl_tac >- ( EVAL_TAC>> fs[ffi_names,LENGTH_data,LENGTH_code,extcalls_def]) \\ strip_tac \\ drule (GEN_ALL target_state_rel_ag32_init) @@ -107,11 +120,12 @@ val cake_startup_clock_def = |> SIMP_RULE bool_ss [GSYM RIGHT_EXISTS_IMP_THM,SKOLEM_THM]); val compile_correct_applied = - MATCH_MP compile_correct_eval cake_compiled - |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO, - with_clos_conf_simp] + MATCH_MP compile_correct_eval (cj 1 compiler32_compiled) + |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm, + GSYM AND_IMP_INTRO,with_clos_conf_simp] |> Q.INST [‘ev’|->‘SOME compiler_instance’] - |> SIMP_RULE (srw_ss()) [add_eval_state_def,opt_eval_config_wf_def,compiler_instance_lemma] + |> SIMP_RULE (srw_ss()) [add_eval_state_def,opt_eval_config_wf_def, + compiler_instance_lemma] |> C MATCH_MP cake_not_fail |> C MATCH_MP ag32_backend_config_ok |> REWRITE_RULE[cake_sem_sing,AND_IMP_INTRO] @@ -131,16 +145,25 @@ Theorem cake_compiled_thm = Theorem cake_installed: SUM (MAP strlen cl) + LENGTH cl ≤ cline_size ∧ LENGTH inp ≤ stdin_size ∧ - is_ag32_init_state (init_memory code data (extcalls config.lab_conf.ffi_names) (cl,inp)) ms0 ⇒ - installed code 0 data 0 config.lab_conf.ffi_names + is_ag32_init_state (init_memory code data + (extcalls info.lab_conf.ffi_names) (cl,inp)) ms0 ⇒ + installed code 0 data 0 info.lab_conf.ffi_names (heap_regs ag32_backend_config.stack_conf.reg_names) - (cake_machine_config) config.lab_conf.shmem_extra + (cake_machine_config) info.lab_conf.shmem_extra (FUNPOW Next (cake_startup_clock ms0 inp cl) ms0) Proof rewrite_tac[ffi_names, extcalls_def, shmem] \\ strip_tac \\ qmatch_asmsub_abbrev_tac ‘init_memory _ _ ff’ - \\ ‘^(ffi_names |> concl |> rand |> rand) = MAP ExtCall ff’ by simp [Abbr‘ff’] + \\ qmatch_goalsub_abbrev_tac ‘installed _ _ _ _ dd’ + \\ ‘dd = SOME (MAP ExtCall ff)’ by + (unabbrev_all_tac + \\ assume_tac (cj 1 compiler32_compiled) + \\ drule ag32_configProofTheory.compile_imp_ffi_names + \\ gvs [compiler32_compiled] + \\ gvs [GSYM compiler32_compiled,ffis] + \\ simp [backendTheory.set_oracle_def, + ag32_configTheory.ag32_backend_config_def]) \\ asm_rewrite_tac [] \\ irule ag32_installed \\ unabbrev_all_tac @@ -200,7 +223,8 @@ Proof QED Theorem FST_ALOOKUP_fastForwardFD_infds: - OPTION_MAP FST (ALOOKUP (fastForwardFD fs fd).infds fd') = OPTION_MAP FST (ALOOKUP fs.infds fd') + OPTION_MAP FST (ALOOKUP (fastForwardFD fs fd).infds fd') = + OPTION_MAP FST (ALOOKUP fs.infds fd') Proof rw[fsFFIPropsTheory.fastForwardFD_def] \\ Cases_on`ALOOKUP fs.infds fd` \\ simp[libTheory.the_def] @@ -212,7 +236,8 @@ Proof QED Theorem FST_ALOOKUP_add_stdo_infds: - OPTION_MAP FST (ALOOKUP (add_stdo fd nm fs out).infds fd') = OPTION_MAP FST (ALOOKUP fs.infds fd') + OPTION_MAP FST (ALOOKUP (add_stdo fd nm fs out).infds fd') = + OPTION_MAP FST (ALOOKUP fs.infds fd') Proof mp_tac TextIOProofTheory.add_stdo_MAP_FST_infds \\ strip_tac @@ -501,7 +526,7 @@ QED Theorem cake_ag32_next: SUM (MAP strlen cl) + LENGTH cl ≤ cline_size ∧ wfcl cl ∧ LENGTH inp ≤ stdin_size ∧ - is_ag32_init_state (init_memory code data (extcalls config.lab_conf.ffi_names) (cl,inp)) ms0 + is_ag32_init_state (init_memory code data (extcalls info.lab_conf.ffi_names) (cl,inp)) ms0 ⇒ ∃k1. ∀k. k1 ≤ k ⇒ let ms = FUNPOW Next k ms0 in diff --git a/compiler/bootstrap/compilation/ag32/32/to_dataBootstrapScript.sml b/compiler/bootstrap/compilation/ag32/32/to_dataBootstrapScript.sml deleted file mode 100644 index 401bd68188..0000000000 --- a/compiler/bootstrap/compilation/ag32/32/to_dataBootstrapScript.sml +++ /dev/null @@ -1,56 +0,0 @@ -(* - Evaluate the 32-bit version of the compiler down to a DataLang - program. -*) -open preamble compiler32ProgTheory - -val _ = new_theory"to_dataBootstrap"; - -(* - Eventually, this file will prove - |- to_data c prog_t1 = ... - |- to_data c prog_t2 = ... - |- ... - where - c = - a default initial config (shared by all targets) - prog_tn = - a prog declaring the entire compiler for target n - - With incremental compilation, we might get away with only one prog, which is - the prog for all the non-target-specific parts of the compiler, but Magnus - suggests incremental compilation like that might be impossible, since some - phases need to know they have the whole program. - - Alternatively, the different to_data theorems for different targets could go - into different theories. The only thing they share is init_conf_def and the - strategy for evaluation. -*) - -val _ = Globals.max_print_depth := 20; - -val cs = compilationLib.compilation_compset(); - -val init_conf_def = zDefine` - init_conf = <| - source_conf := prim_src_config; - clos_conf := clos_to_bvl$default_config - with known_conf := SOME - <| inline_max_body_size := 8; inline_factor := 0; - initial_inline_factor := 0; val_approx_spt := LN |>; - bvl_conf := bvl_to_bvi$default_config with - <| inline_size_limit := 3; exp_cut := 200 |> - |>`; - -(* -val (ls,ty) = compiler32_prog_def |> rconc |> listSyntax.dest_list -val new_prog = listSyntax.mk_list(List.take(ls,80),ty) -val compiler32_prog_thm = mk_thm([],mk_eq(lhs(concl(compiler32_prog_def)),new_prog)); -*) -val compiler32_prog_thm = compiler32_prog_def; - -val to_data_ag32_thm = save_thm("to_data_ag32_thm", - compilationLib.compile_to_data - cs init_conf_def compiler32_prog_thm "data_prog_ag32"); - -val _ = export_theory(); diff --git a/compiler/bootstrap/compilation/ag32/32/to_lab_ag32BootstrapScript.sml b/compiler/bootstrap/compilation/ag32/32/to_lab_ag32BootstrapScript.sml deleted file mode 100644 index dc7e870a29..0000000000 --- a/compiler/bootstrap/compilation/ag32/32/to_lab_ag32BootstrapScript.sml +++ /dev/null @@ -1,60 +0,0 @@ -(* - Evaluate the 32-bit version of the compiler down to a LabLang - program (an assembly program). -*) -open preamble - backendTheory - to_dataBootstrapTheory - ag32_configTheory - ag32_targetTheory - ag32_targetLib asmLib - -val _ = new_theory "to_lab_ag32Bootstrap"; - -val _ = Globals.max_print_depth := 10; - -val new_clos_conf = - (REWRITE_CONV [init_conf_def] THENC EVAL) ``init_conf.clos_conf`` |> concl |> rand - -val bootstrap_conf = - ``((ag32_backend_config - with clos_conf := ^new_clos_conf) - with - bvl_conf updated_by - (λc. c with <| inline_size_limit := 3; exp_cut := 200 |>))`` - -(* coerce the change_config thm to change between configs with the same type - parameter, which in turn coerces to a particular instance of - init_conf/bvi_conf in the to_data theorem, which avoids issues with type - parameters and saving definitions later *) -val change_config_thm = backendTheory.to_data_change_config - |> Q.INST_TYPE [`:'b` |-> `:'a`] - -val to_data_thm0 = - MATCH_MP change_config_thm to_data_ag32_thm - |> Q.GEN`c2` |> ISPEC bootstrap_conf - -val same_config = prove(to_data_thm0 |> concl |> rator |> rand, - REWRITE_TAC[init_conf_def,ag32_backend_config_def] - \\ EVAL_TAC - \\ rw[FLOOKUP_EXT,FUN_EQ_THM,FLOOKUP_UPDATE,FLOOKUP_FUNION] - \\ EVAL_TAC - \\ rpt (IF_CASES_TAC \\ rveq \\ EVAL_TAC)) - -val to_data_thm1 = - MATCH_MP to_data_thm0 same_config - -(* -val (ls,ty) = data_prog_ag32_def |> rconc |> listSyntax.dest_list -val data_prog_ag32' = listSyntax.mk_list(List.take(ls,10),ty) -val data_prog_ag32_shorten = mk_thm([],``data_prog_ag32 = ^data_prog_ag32'``) -val to_data_thm1 = PURE_REWRITE_RULE[data_prog_ag32_shorten]to_data_thm1 -val to_data_thm = to_data_thm1 -*) - -val stack_to_lab_thm = save_thm("stack_to_lab_thm", - compilationLib.compile_to_lab data_prog_ag32_def to_data_thm1 "lab_prog"); - -val () = ml_translatorLib.reset_translation(); - -val () = export_theory(); diff --git a/compiler/bootstrap/compilation/x64/32/Holmakefile b/compiler/bootstrap/compilation/x64/32/Holmakefile index 161c78f6b9..c380805c73 100644 --- a/compiler/bootstrap/compilation/x64/32/Holmakefile +++ b/compiler/bootstrap/compilation/x64/32/Holmakefile @@ -2,7 +2,8 @@ ARCH = x64 WORD_SIZE = 32 INCLUDES = $(CAKEMLDIR)/semantics $(CAKEMLDIR)/basis ../../../..\ ../../../../encoders/asm ../../../../encoders/$(ARCH)\ - ../../../../backend/$(ARCH) ../../../translation + ../../../../backend/$(ARCH) ../../../translation \ + $(CAKEMLDIR)/cv_translator all: $(DEFAULT_TARGETS) README.md cake-$(ARCH)-$(WORD_SIZE).tar.gz .PHONY: all @@ -18,3 +19,7 @@ cake-$(ARCH)-$(WORD_SIZE).tar.gz: cake.S basis_ffi.c Makefile hello.cml how-to.m tar -chzf $@ --transform='s|^|cake-$(ARCH)-$(WORD_SIZE)/|' cake.S basis_ffi.c Makefile hello.cml how-to.md EXTRA_CLEANS = cake.S cake-$(ARCH)-$(WORD_SIZE).tar.gz + +ifdef POLY +HOLHEAP = $(CAKEMLDIR)/cv_translator/cake_compile_heap +endif diff --git a/compiler/bootstrap/compilation/x64/32/README.md b/compiler/bootstrap/compilation/x64/32/README.md index eaee9a89dc..4fc33bbe6e 100644 --- a/compiler/bootstrap/compilation/x64/32/README.md +++ b/compiler/bootstrap/compilation/x64/32/README.md @@ -6,18 +6,12 @@ inside the logic to produce the verified machine code version of the Implements the foreign function interface (FFI) used in the CakeML basis library, as a thin wrapper around the relevant system calls. +[hello.cml](hello.cml): +A simple hello world program in CakeML + [proofs](proofs): This directory contains the end-to-end correctness theorem for the 32-bit version of the CakeML compiler. -[to_dataBootstrapScript.sml](to_dataBootstrapScript.sml): -Evaluate the 32-bit version of the compiler down to a DataLang -program. - -[to_lab_x64BootstrapScript.sml](to_lab_x64BootstrapScript.sml): -Evaluate the 32-bit version of the compiler down to a LabLang -program (an assembly program). - [x64BootstrapScript.sml](x64BootstrapScript.sml): -Evaluate the final part of the 32-bit version of the compiler -into machine code for x64. +Evaluateof the 32-bit version of the compiler into x64 machine code. diff --git a/compiler/bootstrap/compilation/x64/32/proofs/x64BootstrapProofScript.sml b/compiler/bootstrap/compilation/x64/32/proofs/x64BootstrapProofScript.sml index 9ad548108f..a5e486b85b 100644 --- a/compiler/bootstrap/compilation/x64/32/proofs/x64BootstrapProofScript.sml +++ b/compiler/bootstrap/compilation/x64/32/proofs/x64BootstrapProofScript.sml @@ -7,20 +7,22 @@ open preamble val _ = new_theory"x64BootstrapProof"; -val with_clos_conf_simp = prove( - ``(mc_init_ok (x64_backend_config with <| clos_conf := z ; bvl_conf updated_by +Triviality with_clos_conf_simp: + (mc_init_ok (x64_backend_config with <| clos_conf := z ; bvl_conf updated_by (λc. c with <|inline_size_limit := t1; exp_cut := t2|>) |>) = mc_init_ok x64_backend_config) /\ (x.max_app <> 0 /\ (case x.known_conf of NONE => T | SOME k => k.val_approx_spt = LN) ==> (backend_config_ok (x64_backend_config with clos_conf := x) = - backend_config_ok x64_backend_config))``, + backend_config_ok x64_backend_config)) +Proof fs [mc_init_ok_def,FUN_EQ_THM,backend_config_ok_def] - \\ rw [] \\ eq_tac \\ rw [] \\ EVAL_TAC); + \\ rw [] \\ eq_tac \\ rw [] \\ EVAL_TAC +QED Definition compiler_instance_def: compiler_instance = - <| init_state := config_to_inc_config cake_config ; - compiler_fun := compile_inc_progs_for_eval cake_config.lab_conf.asm_conf ; + <| init_state := config_to_inc_config info; + compiler_fun := compile_inc_progs_for_eval x64_config ; config_dom := UNIV ; config_v := BACKEND_INC_CONFIG_v ; decs_dom := decs_allowed ; @@ -29,17 +31,19 @@ End Triviality compiler_instance_lemma: INJ compiler_instance.config_v 𝕌(:inc_config) 𝕌(:semanticPrimitives$v) ∧ - compiler_instance.init_state = config_to_inc_config cake_config ∧ - compiler_instance.compiler_fun = - compile_inc_progs_for_eval cake_config.lab_conf.asm_conf + compiler_instance.init_state = config_to_inc_config info ∧ + compiler_instance.compiler_fun = compile_inc_progs_for_eval x64_config Proof fs [compiler_instance_def] QED -Theorem cake_config_lab_conf_asm_conf: - cake_config.lab_conf.asm_conf = x64_config +Theorem info_asm_conf: + info.lab_conf.asm_conf = x64_config Proof - once_rewrite_tac [cake_config_def] \\ EVAL_TAC + assume_tac $ cj 1 compiler32_compiled + \\ drule compile_asm_config_eq + \\ gvs [backendTheory.set_oracle_def] + \\ strip_tac \\ EVAL_TAC QED val cake_io_events_def = new_specification("cake_io_events_def",["cake_io_events"], @@ -55,11 +59,12 @@ val (cake_sem,cake_output) = cake_io_events_def |> SPEC_ALL |> UNDISCH |> CONJ_P val (cake_not_fail,cake_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail cake_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct_eval cake_compiled - |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO, - with_clos_conf_simp] + MATCH_MP compile_correct_eval (cj 1 compiler32_compiled) + |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm, + GSYM AND_IMP_INTRO,with_clos_conf_simp] |> Q.INST [‘ev’|->‘SOME compiler_instance’] - |> SIMP_RULE (srw_ss()) [add_eval_state_def,opt_eval_config_wf_def,compiler_instance_lemma] + |> SIMP_RULE (srw_ss()) [add_eval_state_def,opt_eval_config_wf_def, + compiler_instance_lemma,info_asm_conf] |> C MATCH_MP cake_not_fail |> C MATCH_MP x64_backend_config_ok |> REWRITE_RULE[cake_sem_sing,AND_IMP_INTRO] diff --git a/compiler/bootstrap/compilation/x64/32/to_dataBootstrapScript.sml b/compiler/bootstrap/compilation/x64/32/to_dataBootstrapScript.sml deleted file mode 100644 index 688f989789..0000000000 --- a/compiler/bootstrap/compilation/x64/32/to_dataBootstrapScript.sml +++ /dev/null @@ -1,56 +0,0 @@ -(* - Evaluate the 32-bit version of the compiler down to a DataLang - program. -*) -open preamble compiler32ProgTheory - -val _ = new_theory"to_dataBootstrap"; - -(* - Eventually, this file will prove - |- to_data c prog_t1 = ... - |- to_data c prog_t2 = ... - |- ... - where - c = - a default initial config (shared by all targets) - prog_tn = - a prog declaring the entire compiler for target n - - With incremental compilation, we might get away with only one prog, which is - the prog for all the non-target-specific parts of the compiler, but Magnus - suggests incremental compilation like that might be impossible, since some - phases need to know they have the whole program. - - Alternatively, the different to_data theorems for different targets could go - into different theories. The only thing they share is init_conf_def and the - strategy for evaluation. -*) - -val _ = Globals.max_print_depth := 20; - -val cs = compilationLib.compilation_compset(); - -val init_conf_def = zDefine` - init_conf = <| - source_conf := prim_src_config; - clos_conf := clos_to_bvl$default_config - with known_conf := SOME - <| inline_max_body_size := 8; inline_factor := 0; - initial_inline_factor := 0; val_approx_spt := LN |>; - bvl_conf := bvl_to_bvi$default_config with - <| inline_size_limit := 3; exp_cut := 200 |> - |>`; - -(* -val (ls,ty) = compiler_prog_def |> rconc |> listSyntax.dest_list -val new_prog = listSyntax.mk_list(List.take(ls,80),ty) -val compiler_prog_thm = mk_thm([],mk_eq(lhs(concl(compiler_prog_def)),new_prog)); -*) -val compiler32_prog_thm = compiler32_prog_def; - -val to_data_x64_thm = save_thm("to_data_x64_thm", - compilationLib.compile_to_data - cs init_conf_def compiler32_prog_thm "data_prog_x64"); - -val _ = export_theory(); diff --git a/compiler/bootstrap/compilation/x64/32/to_lab_x64BootstrapScript.sml b/compiler/bootstrap/compilation/x64/32/to_lab_x64BootstrapScript.sml deleted file mode 100644 index d533075257..0000000000 --- a/compiler/bootstrap/compilation/x64/32/to_lab_x64BootstrapScript.sml +++ /dev/null @@ -1,60 +0,0 @@ -(* - Evaluate the 32-bit version of the compiler down to a LabLang - program (an assembly program). -*) -open preamble - backendTheory - to_dataBootstrapTheory - x64_configTheory - x64_targetTheory - x64_targetLib asmLib - -val _ = new_theory "to_lab_x64Bootstrap"; - -val _ = Globals.max_print_depth := 10; - -val new_clos_conf = - (REWRITE_CONV [init_conf_def] THENC EVAL) ``init_conf.clos_conf`` |> concl |> rand - -val bootstrap_conf = - ``((x64_backend_config - with clos_conf := ^new_clos_conf) - with - bvl_conf updated_by - (λc. c with <| inline_size_limit := 3; exp_cut := 200 |>))`` - -(* coerce the change_config thm to change between configs with the same type - parameter, which in turn coerces to a particular instance of - init_conf/bvi_conf in the to_data theorem, which avoids issues with type - parameters and saving definitions later *) -val change_config_thm = backendTheory.to_data_change_config - |> Q.INST_TYPE [`:'b` |-> `:'a`] - -val to_data_thm0 = - MATCH_MP change_config_thm to_data_x64_thm - |> Q.GEN`c2` |> ISPEC bootstrap_conf - -val same_config = prove(to_data_thm0 |> concl |> rator |> rand, - REWRITE_TAC[init_conf_def,x64_backend_config_def] - \\ EVAL_TAC - \\ rw[FLOOKUP_EXT,FUN_EQ_THM,FLOOKUP_UPDATE,FLOOKUP_FUNION] - \\ EVAL_TAC - \\ rpt (IF_CASES_TAC \\ rveq \\ EVAL_TAC)) - -val to_data_thm1 = - MATCH_MP to_data_thm0 same_config - -(* -val (ls,ty) = data_prog_x64_def |> rconc |> listSyntax.dest_list -val data_prog_x64' = listSyntax.mk_list(List.take(ls,10),ty) -val data_prog_x64_shorten = mk_thm([],``data_prog_x64 = ^data_prog_x64'``) -val to_data_thm1 = PURE_REWRITE_RULE[data_prog_x64_shorten]to_data_thm1 -val to_data_thm = to_data_thm1 -*) - -val stack_to_lab_thm = save_thm("stack_to_lab_thm", - compilationLib.compile_to_lab data_prog_x64_def to_data_thm1 "lab_prog"); - -val () = ml_translatorLib.reset_translation(); - -val () = export_theory(); diff --git a/compiler/bootstrap/compilation/x64/32/x64BootstrapScript.sml b/compiler/bootstrap/compilation/x64/32/x64BootstrapScript.sml index 0aed637211..a3e2625dac 100644 --- a/compiler/bootstrap/compilation/x64/32/x64BootstrapScript.sml +++ b/compiler/bootstrap/compilation/x64/32/x64BootstrapScript.sml @@ -1,292 +1,11 @@ (* - Evaluate the final part of the 32-bit version of the compiler - into machine code for x64. + Evaluateof the 32-bit version of the compiler into x64 machine code. *) -open preamble to_lab_x64BootstrapTheory compilationLib +open preamble compiler32ProgTheory eval_cake_compile_x64Lib; -val _ = new_theory "x64Bootstrap"; +val _ = new_theory "x64Bootstrap" -val () = ml_translatorLib.reset_translation(); +Theorem compiler32_compiled = + eval_cake_compile_x64 "" compiler32_prog_def "cake.S"; -(* - val lab_prog_def = - let - val (ls,ty) = to_lab_x64BootstrapTheory.lab_prog_def |> rconc |> listSyntax.dest_list - val ls' = listSyntax.mk_list(List.take(List.drop(ls,2000),40),ty) - in mk_thm([],``lab_prog = ^ls'``) end -*) - -val filename = "cake.S" - -val bootstrap_thm = - compilationLib.cbv_to_bytes_x64 - stack_to_lab_thm lab_prog_def - "cake_code" "cake_data" "cake_config" filename; - -val cake_compiled = save_thm("cake_compiled", bootstrap_thm); - -(* avoid saving the long list of bytes in the Theory.sml file - they can only be found in the exported cake.S file *) -val _ = Theory.delete_binding "cake_code_def"; - -val _ = export_theory(); - -(* -val Label_tm = ela2 |> SPEC_ALL |> concl |> lhs |> rator |> rand |> rator |> rand |> repeat rator; -val Asm_tm = ela3 |> SPEC_ALL |> concl |> lhs |> rator |> rand |> rator |> rand |> repeat rator; -val LabAsm_tm = ela4 |> SPEC_ALL |> concl |> lhs |> rator |> rand |> rator |> rand |> repeat rator; - -fun enc_lines_again_rule labs_def = -let - fun f th = let - val ls = th |> rconc |> rator |> rand - in if listSyntax.is_nil ls then - CONV_RULE(RAND_CONV (REWR_CONV ela1 THENC RATOR_CONV(RAND_CONV eval))) th - else let - val tm = listSyntax.dest_cons ls |> #1 |> repeat rator - val th = - if same_const Label_tm tm then - CONV_RULE(RAND_CONV (REWR_CONV ela2 THENC add_pos_conv)) th - else if same_const Asm_tm tm then - CONV_RULE(RAND_CONV (REWR_CONV ela3 THENC add_pos_conv)) th - else - let - val _ = assert (same_const LabAsm_tm) tm - val th = CONV_RULE(RAND_CONV ( - REWR_CONV ela4 THENC - RAND_CONV (RATOR_CONV(RAND_CONV(REWR_CONV labs_def)) THENC eval) THENC - REWR_CONV LET_THM THENC BETA_CONV THENC - RATOR_CONV(RATOR_CONV(RAND_CONV eval)))) th - val tm = th |> rconc |> rator |> rator |> rand - in - if same_const T tm then - CONV_RULE(RAND_CONV (REWR_CONV COND_T THENC add_pos_conv)) th - else (assert (same_const F) tm; - CONV_RULE(RAND_CONV (REWR_CONV COND_F THENC - RAND_CONV eval THENC REWR_CONV LET_THM THENC BETA_CONV THENC - RAND_CONV eval THENC REWR_CONV LET_THM THENC BETA_CONV THENC - add_pos_conv THENC - RAND_CONV(RAND_CONV(RAND_CONV eval THENC REWR_CONV AND_T)))) th) - end - in - f th - end - end -in f end -*) - -(* -fun enc_lines_again_rule labs_def = -let - fun f th = - let - val (th,b) = - (* Asm *) - (CONV_RULE(RAND_CONV (REWR_CONV ela3 THENC add_pos_conv)) th,true) - handle HOL_ERR _ => - (* Label *) - (CONV_RULE(RAND_CONV (REWR_CONV ela2 THENC add_pos_conv)) th,true) - handle HOL_ERR _ => - (* LabAsm *) - let - val th = CONV_RULE(RAND_CONV ( - REWR_CONV ela4 THENC - RAND_CONV (RATOR_CONV(RAND_CONV(REWR_CONV labs_def)) THENC eval) THENC - let_CONV THENC - RATOR_CONV(RATOR_CONV(RAND_CONV eval)))) th - in - (* no offset update *) - (CONV_RULE(RAND_CONV (REWR_CONV COND_T THENC add_pos_conv)) th,true) - handle HOL_ERR _ => - (* offset update *) - (CONV_RULE(RAND_CONV (REWR_CONV COND_F THENC - RAND_CONV eval THENC let_CONV THENC - RAND_CONV eval THENC let_CONV THENC - add_pos_conv THENC - RAND_CONV(RAND_CONV(RAND_CONV eval THENC REWR_CONV AND_T)))) th,true) - end - handle HOL_ERR _ => - (* nil *) - (CONV_RULE(RAND_CONV (REWR_CONV ela1 THENC RATOR_CONV(RAND_CONV eval))) th,false) - in if b then f th else th end -in f end - -fun enc_lines_again_conv labs_def = enc_lines_again_rule labs_def o REFL -*) - -(* -fun enc_lines_again_conv labs_def tm = tm |> ( - IFC - ((REWR_CONV ela3 THENC add_pos_conv) ORELSEC - (REWR_CONV ela2 THENC add_pos_conv) ORELSEC - (REWR_CONV ela4 THENC - RAND_CONV (RATOR_CONV(RAND_CONV(REWR_CONV labs_def)) THENC eval) THENC - let_CONV THENC - RATOR_CONV(RATOR_CONV(RAND_CONV eval)) THENC - ((REWR_CONV COND_T THENC add_pos_conv) ORELSEC - (REWR_CONV COND_F THENC - RAND_CONV eval THENC let_CONV THENC - RAND_CONV eval THENC let_CONV THENC - add_pos_conv THENC - RAND_CONV(RAND_CONV(RAND_CONV eval THENC REWR_CONV AND_T)))))) - (enc_lines_again_conv labs_def) - (REWR_CONV ela1 THENC RATOR_CONV(RAND_CONV eval))) -*) - -(* -fun enc_lines_again_conv labs_def tm = tm |> ( - IFC - ((REWR_CONV ela3) ORELSEC - (REWR_CONV ela2) ORELSEC - (REWR_CONV ela4 THENC - RAND_CONV (RATOR_CONV(RAND_CONV(REWR_CONV labs_def)) THENC eval) THENC - let_CONV THENC - RATOR_CONV(RATOR_CONV(RAND_CONV eval)) THENC - ((REWR_CONV COND_T) ORELSEC - (REWR_CONV COND_F THENC - RAND_CONV eval THENC let_CONV THENC - RAND_CONV eval THENC let_CONV)))) - (enc_lines_again_conv labs_def) - (REWR_CONV ela1 THENC eval)) -*) - -(* -fun enc_lines_again_conv tm = tm |> ( - IFC - ((REWR_CONV ela3) ORELSEC - (REWR_CONV ela2) ORELSEC - (REWR_CONV ela4 THENC - RAND_CONV eval THENC let_CONV THENC - RATOR_CONV(RATOR_CONV(RAND_CONV eval)) THENC - ((REWR_CONV COND_T) ORELSEC - (REWR_CONV COND_F THENC - RAND_CONV eval THENC let_CONV THENC - RAND_CONV eval THENC let_CONV)))) - (enc_lines_again_conv) - (REWR_CONV ela1 THENC eval)) -*) - -(* -fun enc_lines_again_conv tm = tm |> ( - IFC - ((REWR_CONV ela3) ORELSEC - (REWR_CONV ela2) ORELSEC - (REWR_CONV ela4 THENC - RAND_CONV eval THENC let_CONV THENC - RATOR_CONV(RATOR_CONV(RAND_CONV eval)) THENC - COND_CONV THENC - TRY_CONV - (RAND_CONV eval THENC let_CONV THENC - RAND_CONV eval THENC let_CONV))) - (enc_lines_again_conv) - (REWR_CONV ela1 THENC eval)) -*) - -(* -fun enc_lines_again_conv tm = tm |> ( - IFC - ((REWR_CONV ela3) ORELSEC - (REWR_CONV ela2) ORELSEC - (REWR_CONV ela4 THENC - RAND_CONV eval THENC let_CONV THENC - RATOR_CONV(RATOR_CONV(RAND_CONV eval)) THENC - COND_CONV THENC REPEATC let_CONV)) - (enc_lines_again_conv) - (REWR_CONV ela1 THENC eval)) -*) - -(* -val encoded_prog_thm_prefix = - let - val (car,cdr) = dest_comb (concl encoded_prog_thm) - val (ls,ty) = listSyntax.dest_list cdr - val ls' = List.take(ls,10) - in mk_thm([],mk_comb(car,listSyntax.mk_list(ls',ty))) end -val encoded_prog_defs_prefix = - List.take(List.rev encoded_prog_defs,10) - -(* rule using exceptions *) -val enc_secs_again_thm = - tm13 |> timez "enc_secs_again" ( - RAND_CONV(REWR_CONV encoded_prog_thm_prefix) THENC - enc_secs_again_conv - "enc_again_" (enc_lines_again_conv computed_labs_def) 0 - encoded_prog_defs_prefix) -2m15s - -(* custom conv *) -val enc_secs_again_thm = - tm13 |> timez "enc_secs_again" ( - RAND_CONV(REWR_CONV encoded_prog_thm_prefix) THENC - enc_secs_again_conv - "enc_again_" (enc_lines_again_conv computed_labs_def) 0 - encoded_prog_defs_prefix) -2m17s - -(* custom conv with IFC *) -val enc_secs_again_thm = - tm13 |> timez "enc_secs_again" ( - RAND_CONV(REWR_CONV encoded_prog_thm_prefix) THENC - enc_secs_again_conv - "enc_again_" (enc_lines_again_conv computed_labs_def) 0 - encoded_prog_defs_prefix) -2m18s - -val enc_secs_again_thm = - tm13 |> timez "enc_secs_again" ( - RAND_CONV(REWR_CONV encoded_prog_thm_prefix) THENC - enc_secs_again_conv - "enc_again_" - (PATH_CONV "llllr" (REWR_CONV computed_labs_def) THENC eval) - 0 - encoded_prog_defs_prefix) -14.4s - -val (dth::dths) = encoded_prog_defs_prefix -val tm = tm13 |> RAND_CONV(REWR_CONV encoded_prog_thm_prefix) |> rconc -val th1 = tm |> (REWR_CONV esc THENC (RAND_CONV(RATOR_CONV(RAND_CONV(REWR_CONV dth))))) -val tm = th1 |> rconc |> rand - -val test1 = - time (PATH_CONV "llllr" (REWR_CONV computed_labs_def) THENC eval) tm -1.1s - -val th = REFL tm - -(* rule with exceptions *) -val test2 = - time (enc_lines_again_conv computed_labs_def) tm -17.1s -val check = rconc test1 = rconc test2 - -(* custom conv with IFC *) -val test3 = - time (enc_lines_again_conv computed_labs_def) tm -30.6s -val check = rconc test1 = rconc test3 - -(* custom conv with IFC and delayed leaves *) -val test4 = - time (enc_lines_again_conv computed_labs_def) tm -31.3s -val check = rconc test1 = rconc test4 - -(* custom conv with labs expanded *) -val test5 = - time (PATH_CONV "llllr" (REWR_CONV computed_labs_def) THENC enc_lines_again_conv) tm -21.2s -val check = rconc test1 = rconc test5 - -(* custom conv with labs expanded using COND_CONV *) -val test6 = - time (PATH_CONV "llllr" (REWR_CONV computed_labs_def) THENC enc_lines_again_conv) tm -19.4s -val check = rconc test1 = rconc test6 - -(* custom conv with labs expanded using COND_CONV and deferring all computations *) -val test7 = - time (PATH_CONV "llllr" (REWR_CONV computed_labs_def) THENC enc_lines_again_conv) tm -19.6s -val check = rconc test1 = rconc test7 - -*) +val _ = export_theory (); diff --git a/compiler/bootstrap/compilation/x64/64/Holmakefile b/compiler/bootstrap/compilation/x64/64/Holmakefile index 8ab472bff8..926c6d8bc5 100644 --- a/compiler/bootstrap/compilation/x64/64/Holmakefile +++ b/compiler/bootstrap/compilation/x64/64/Holmakefile @@ -2,7 +2,8 @@ ARCH = x64 WORD_SIZE = 64 INCLUDES = $(CAKEMLDIR)/semantics $(CAKEMLDIR)/basis ../../../.. $(CAKEMLDIR)/unverified/sexpr-bootstrap \ ../../../../encoders/asm ../../../../encoders/$(ARCH)\ - ../../../../backend/$(ARCH) ../../../../backend/serialiser ../../../translation + ../../../../backend/$(ARCH) ../../../../backend/serialiser ../../../translation \ + $(CAKEMLDIR)/cv_translator all: $(DEFAULT_TARGETS) README.md cake-$(ARCH)-$(WORD_SIZE).tar.gz .PHONY: all @@ -13,7 +14,7 @@ README.md: $(CAKEMLDIR)/developers/readme_gen readmePrefix $(patsubst %,%readmeP $(protect $(CAKEMLDIR)/developers/readme_gen) $(README_SOURCES) cake.S: *$(ARCH)BootstrapScript.sml -config_enc_str.txt: *$(ARCH)_config_encScript.sml +config_enc_str.txt: *$(ARCH)BootstrapScript.sml cake-sexpr-32: *sexprBootstrap32Script.sml cake-sexpr-64: *sexprBootstrap64Script.sml @@ -26,3 +27,7 @@ cake-$(ARCH)-$(WORD_SIZE).tar.gz: cake.S basis_ffi.c Makefile hello.cml how-to.m diff output expected_output # returns non-zero if files differ EXTRA_CLEANS = cake.S cake-$(ARCH)-$(WORD_SIZE).tar.gz cake test-hello.cake output expected_output + +ifdef POLY +HOLHEAP = $(CAKEMLDIR)/cv_translator/cake_compile_heap +endif diff --git a/compiler/bootstrap/compilation/x64/64/README.md b/compiler/bootstrap/compilation/x64/64/README.md index e09bce30bb..de5f354737 100644 --- a/compiler/bootstrap/compilation/x64/64/README.md +++ b/compiler/bootstrap/compilation/x64/64/README.md @@ -29,17 +29,5 @@ definition for the 64-bit version of the compiler. A hello world program used for testing that the bootstrapped compiler was built succesfully. -[to_dataBootstrapScript.sml](to_dataBootstrapScript.sml): -Evaluate the 32-bit version of the compiler down to a DataLang -program. - -[to_lab_x64BootstrapScript.sml](to_lab_x64BootstrapScript.sml): -Evaluate the 64-bit version of the compiler down to a LabLang -program (an assembly program). - [x64BootstrapScript.sml](x64BootstrapScript.sml): -Evaluate the final part of the 64-bit version of the compiler -into machine code for x64. - -[x64_config_encScript.sml](x64_config_encScript.sml): -Encoding of compiler state +Evaluate the 64-bit version of the compiler into x64 machine code. diff --git a/compiler/bootstrap/compilation/x64/64/proofs/x64BootstrapProofScript.sml b/compiler/bootstrap/compilation/x64/64/proofs/x64BootstrapProofScript.sml index 9c3e631062..7f49f8dac7 100644 --- a/compiler/bootstrap/compilation/x64/64/proofs/x64BootstrapProofScript.sml +++ b/compiler/bootstrap/compilation/x64/64/proofs/x64BootstrapProofScript.sml @@ -8,20 +8,22 @@ open preamble val _ = new_theory"x64BootstrapProof"; -val with_clos_conf_simp = prove( - ``(mc_init_ok (x64_backend_config with <| clos_conf := z ; bvl_conf updated_by +Triviality with_clos_conf_simp: + (mc_init_ok (x64_backend_config with <| clos_conf := z ; bvl_conf updated_by (λc. c with <|inline_size_limit := t1; exp_cut := t2|>) |>) = mc_init_ok x64_backend_config) /\ (x.max_app <> 0 /\ (case x.known_conf of NONE => T | SOME k => k.val_approx_spt = LN) ==> (backend_config_ok (x64_backend_config with clos_conf := x) = - backend_config_ok x64_backend_config))``, + backend_config_ok x64_backend_config)) +Proof fs [mc_init_ok_def,FUN_EQ_THM,backend_config_ok_def] - \\ rw [] \\ eq_tac \\ rw [] \\ EVAL_TAC); + \\ rw [] \\ eq_tac \\ rw [] \\ EVAL_TAC +QED Definition compiler_instance_def: compiler_instance = - <| init_state := config_to_inc_config cake_config ; - compiler_fun := compile_inc_progs_for_eval cake_config.lab_conf.asm_conf ; + <| init_state := config_to_inc_config info; + compiler_fun := compile_inc_progs_for_eval x64_config ; config_dom := UNIV ; config_v := BACKEND_INC_CONFIG_v ; decs_dom := decs_allowed ; @@ -30,17 +32,19 @@ End Triviality compiler_instance_lemma: INJ compiler_instance.config_v 𝕌(:inc_config) 𝕌(:semanticPrimitives$v) ∧ - compiler_instance.init_state = config_to_inc_config cake_config ∧ - compiler_instance.compiler_fun = - compile_inc_progs_for_eval cake_config.lab_conf.asm_conf + compiler_instance.init_state = config_to_inc_config info ∧ + compiler_instance.compiler_fun = compile_inc_progs_for_eval x64_config Proof fs [compiler_instance_def] QED -Theorem cake_config_lab_conf_asm_conf: - cake_config.lab_conf.asm_conf = x64_config +Theorem info_asm_conf: + info.lab_conf.asm_conf = x64_config Proof - once_rewrite_tac [cake_config_def] \\ EVAL_TAC + assume_tac $ cj 1 compiler64_compiled + \\ drule compile_asm_config_eq + \\ gvs [backendTheory.set_oracle_def] + \\ strip_tac \\ EVAL_TAC QED val cake_io_events_def = new_specification("cake_io_events_def",["cake_io_events"], @@ -56,11 +60,12 @@ val (cake_sem,cake_output) = cake_io_events_def |> SPEC_ALL |> UNDISCH |> CONJ_P val (cake_not_fail,cake_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail cake_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct_eval cake_compiled - |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO, - with_clos_conf_simp] + MATCH_MP compile_correct_eval (cj 1 compiler64_compiled) + |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm, + GSYM AND_IMP_INTRO,with_clos_conf_simp] |> Q.INST [‘ev’|->‘SOME compiler_instance’] - |> SIMP_RULE (srw_ss()) [add_eval_state_def,opt_eval_config_wf_def,compiler_instance_lemma] + |> SIMP_RULE (srw_ss()) [add_eval_state_def,opt_eval_config_wf_def, + compiler_instance_lemma,info_asm_conf] |> C MATCH_MP cake_not_fail |> C MATCH_MP x64_backend_config_ok |> REWRITE_RULE[cake_sem_sing,AND_IMP_INTRO] @@ -87,9 +92,7 @@ Theorem mk_init_eval_state_lemma = source_evalProofTheory.mk_init_eval_state_def] |> ONCE_REWRITE_RULE [mk_compiler_fun_from_ci_tuple] |> SIMP_RULE (srw_ss()) [source_evalProofTheory.mk_compiler_fun_from_ci_def, - GSYM compiler_inst_def,cake_config_lab_conf_asm_conf]; - -Overload config_env_str = “encode_backend_config (config_to_inc_config cake_config)” + GSYM compiler_inst_def]; Overload init_eval_state_for = “λcl fs. (init_state (basis_ffi cl fs) with @@ -116,24 +119,25 @@ Theorem repl_not_fail = |> Q.GEN ‘s’ |> ISPEC (mk_init_eval_state_lemma |> concl |> rand |> rand) |> REWRITE_RULE [GSYM mk_init_eval_state_lemma] |> SIMP_RULE (srw_ss()) [IN_DEF] - |> Q.INST [‘conf’|->‘config_to_inc_config cake_config’] + |> Q.GEN ‘conf’ |> Q.SPEC ‘config_to_inc_config info’ |> REWRITE_RULE [GSYM (SIMP_CONV (srw_ss()) [] “hasFreeFD fs”)]; Overload basis_init_ok = “λcl fs. STD_streams fs ∧ wfFS fs ∧ wfcl cl ∧ hasFreeFD fs ∧ - file_content fs «config_enc_str.txt» = SOME config_env_str”; + file_content fs «config_enc_str.txt» = SOME conf”; Theorem repl_not_fail_thm: has_repl_flag (TL cl) ∧ basis_init_ok cl fs ⇒ Fail ∉ semantics_prog (init_eval_state_for cl fs) init_env compiler64_prog Proof rw [IN_DEF] \\ irule repl_not_fail \\ fs [] + \\ simp [compiler64_compiled] QED val compile_correct_applied2 = - MATCH_MP compile_correct_eval cake_compiled - |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO, - with_clos_conf_simp] + MATCH_MP compile_correct_eval (cj 1 compiler64_compiled) + |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm, + GSYM AND_IMP_INTRO,with_clos_conf_simp] |> Q.INST [‘ev’|->‘SOME compiler_instance’] |> SIMP_RULE (srw_ss()) [add_eval_state_def,opt_eval_config_wf_def, x64_configProofTheory.x64_backend_config_ok,compiler_instance_lemma] @@ -144,12 +148,12 @@ Definition repl_ready_to_run_def: ∃cbspace data_sp. has_repl_flag (TL cl) ∧ wfcl cl ∧ wfFS fs ∧ STD_streams fs ∧ hasFreeFD fs ∧ - file_content fs «config_enc_str.txt» = SOME config_env_str ∧ + file_content fs «config_enc_str.txt» = SOME conf ∧ mc_conf_ok mc ∧ mc_init_ok x64_backend_config mc ∧ - installed cake_code cbspace cake_data data_sp - cake_config.lab_conf.ffi_names + installed code cbspace data data_sp + info.lab_conf.ffi_names (heap_regs x64_backend_config.stack_conf.reg_names) - mc cake_config.lab_conf.shmem_extra ms + mc info.lab_conf.shmem_extra ms End Overload machine_sem = “λffi (mc,ms). targetSem$machine_sem mc ffi ms” @@ -164,6 +168,7 @@ Proof PairCases_on ‘ms’ \\ rw [IN_DEF,repl_ready_to_run_def] \\ irule compile_correct_applied2 \\ fs [] \\ first_x_assum $ irule_at Any + \\ rewrite_tac [info_asm_conf] QED Triviality isPREFIX_MEM: diff --git a/compiler/bootstrap/compilation/x64/64/to_dataBootstrapScript.sml b/compiler/bootstrap/compilation/x64/64/to_dataBootstrapScript.sml deleted file mode 100644 index e1293b78c0..0000000000 --- a/compiler/bootstrap/compilation/x64/64/to_dataBootstrapScript.sml +++ /dev/null @@ -1,59 +0,0 @@ -(* - Evaluate the 32-bit version of the compiler down to a DataLang - program. -*) -open preamble compiler64ProgTheory - -val _ = new_theory"to_dataBootstrap"; - -(* - Eventually, this file will prove - |- to_data c prog_t1 = ... - |- to_data c prog_t2 = ... - |- ... - where - c = - a default initial config (shared by all targets) - prog_tn = - a prog declaring the entire compiler for target n - - With incremental compilation, we might get away with only one prog, which is - the prog for all the non-target-specific parts of the compiler, but Magnus - suggests incremental compilation like that might be impossible, since some - phases need to know they have the whole program. - - Alternatively, the different to_data theorems for different targets could go - into different theories. The only thing they share is init_conf_def and the - strategy for evaluation. -*) - -val _ = Globals.max_print_depth := 20; - -val cs = compilationLib.compilation_compset(); - -val init_conf_def = zDefine` - init_conf = <| - source_conf := prim_src_config; - clos_conf := clos_to_bvl$default_config - with known_conf := SOME - <| inline_max_body_size := 8; inline_factor := 0; - initial_inline_factor := 0; val_approx_spt := LN |>; - bvl_conf := bvl_to_bvi$default_config with - <| inline_size_limit := 3; exp_cut := 200 |> - |>`; - -(* -val (ls,ty) = compiler_prog_def |> rconc |> listSyntax.dest_list -val new_prog = listSyntax.mk_list(List.take(ls,80),ty) -val compiler_prog_thm = mk_thm([],mk_eq(lhs(concl(compiler_prog_def)),new_prog)); -*) -val compiler64_prog_thm = compiler64_prog_def; - -(* uncomment the line below for debugging purposes *) -(* val compiler64_prog_thm = prove(“compiler64_prog = []”,cheat); *) - -val to_data_x64_thm = save_thm("to_data_x64_thm", - compilationLib.compile_to_data - cs init_conf_def compiler64_prog_thm "data_prog_x64"); - -val _ = export_theory(); diff --git a/compiler/bootstrap/compilation/x64/64/to_lab_x64BootstrapScript.sml b/compiler/bootstrap/compilation/x64/64/to_lab_x64BootstrapScript.sml deleted file mode 100644 index a8bb2b3904..0000000000 --- a/compiler/bootstrap/compilation/x64/64/to_lab_x64BootstrapScript.sml +++ /dev/null @@ -1,60 +0,0 @@ -(* - Evaluate the 64-bit version of the compiler down to a LabLang - program (an assembly program). -*) -open preamble - backendTheory - to_dataBootstrapTheory - x64_configTheory - x64_targetTheory - x64_targetLib asmLib - -val _ = new_theory "to_lab_x64Bootstrap"; - -val _ = Globals.max_print_depth := 10; - -val new_clos_conf = - (REWRITE_CONV [init_conf_def] THENC EVAL) ``init_conf.clos_conf`` |> concl |> rand - -val bootstrap_conf = - ``((x64_backend_config - with clos_conf := ^new_clos_conf) - with - bvl_conf updated_by - (λc. c with <| inline_size_limit := 3; exp_cut := 200 |>))`` - -(* coerce the change_config thm to change between configs with the same type - parameter, which in turn coerces to a particular instance of - init_conf/bvi_conf in the to_data theorem, which avoids issues with type - parameters and saving definitions later *) -val change_config_thm = backendTheory.to_data_change_config - |> Q.INST_TYPE [`:'b` |-> `:'a`] - -val to_data_thm0 = - MATCH_MP change_config_thm to_data_x64_thm - |> Q.GEN`c2` |> ISPEC bootstrap_conf - -val same_config = prove(to_data_thm0 |> concl |> rator |> rand, - REWRITE_TAC[init_conf_def,x64_backend_config_def] - \\ EVAL_TAC - \\ rw[FLOOKUP_EXT,FUN_EQ_THM,FLOOKUP_UPDATE,FLOOKUP_FUNION] - \\ EVAL_TAC - \\ rpt (IF_CASES_TAC \\ rveq \\ EVAL_TAC)) - -val to_data_thm1 = - MATCH_MP to_data_thm0 same_config - -(* -val (ls,ty) = data_prog_x64_def |> rconc |> listSyntax.dest_list -val data_prog_x64' = listSyntax.mk_list(List.take(ls,10),ty) -val data_prog_x64_shorten = mk_thm([],``data_prog_x64 = ^data_prog_x64'``) -val to_data_thm1 = PURE_REWRITE_RULE[data_prog_x64_shorten]to_data_thm1 -val to_data_thm = to_data_thm1 -*) - -val stack_to_lab_thm = save_thm("stack_to_lab_thm", - compilationLib.compile_to_lab data_prog_x64_def to_data_thm1 "lab_prog"); - -val () = ml_translatorLib.reset_translation(); - -val () = export_theory(); diff --git a/compiler/bootstrap/compilation/x64/64/x64BootstrapScript.sml b/compiler/bootstrap/compilation/x64/64/x64BootstrapScript.sml index 841c449164..0ce558e91d 100644 --- a/compiler/bootstrap/compilation/x64/64/x64BootstrapScript.sml +++ b/compiler/bootstrap/compilation/x64/64/x64BootstrapScript.sml @@ -1,292 +1,16 @@ (* - Evaluate the final part of the 64-bit version of the compiler - into machine code for x64. + Evaluate the 64-bit version of the compiler into x64 machine code. *) -open preamble to_lab_x64BootstrapTheory compilationLib +open preamble compiler64ProgTheory eval_cake_compile_x64Lib; -val _ = new_theory "x64Bootstrap"; +val _ = new_theory "x64Bootstrap" -val () = ml_translatorLib.reset_translation(); +Theorem compiler64_compiled = + eval_cake_compile_x64_general + { prefix = "" + , conf_def = x64_configTheory.x64_backend_config_def + , prog_def = compiler64_prog_def + , output_filename = "cake.S" + , output_conf_filename = SOME "config_enc_str.txt" }; -(* - val lab_prog_def = - let - val (ls,ty) = to_lab_x64BootstrapTheory.lab_prog_def |> rconc |> listSyntax.dest_list - val ls' = listSyntax.mk_list(List.take(List.drop(ls,2000),40),ty) - in mk_thm([],``lab_prog = ^ls'``) end -*) - -val filename = "cake.S" - -val bootstrap_thm = - compilationLib.cbv_to_bytes_x64 - stack_to_lab_thm lab_prog_def - "cake_code" "cake_data" "cake_config" filename; - -val cake_compiled = save_thm("cake_compiled", bootstrap_thm); - -(* avoid saving the long list of bytes in the Theory.sml file - they can only be found in the exported cake.S file *) -val _ = Theory.delete_binding "cake_code_def"; - -val _ = export_theory(); - -(* -val Label_tm = ela2 |> SPEC_ALL |> concl |> lhs |> rator |> rand |> rator |> rand |> repeat rator; -val Asm_tm = ela3 |> SPEC_ALL |> concl |> lhs |> rator |> rand |> rator |> rand |> repeat rator; -val LabAsm_tm = ela4 |> SPEC_ALL |> concl |> lhs |> rator |> rand |> rator |> rand |> repeat rator; - -fun enc_lines_again_rule labs_def = -let - fun f th = let - val ls = th |> rconc |> rator |> rand - in if listSyntax.is_nil ls then - CONV_RULE(RAND_CONV (REWR_CONV ela1 THENC RATOR_CONV(RAND_CONV eval))) th - else let - val tm = listSyntax.dest_cons ls |> #1 |> repeat rator - val th = - if same_const Label_tm tm then - CONV_RULE(RAND_CONV (REWR_CONV ela2 THENC add_pos_conv)) th - else if same_const Asm_tm tm then - CONV_RULE(RAND_CONV (REWR_CONV ela3 THENC add_pos_conv)) th - else - let - val _ = assert (same_const LabAsm_tm) tm - val th = CONV_RULE(RAND_CONV ( - REWR_CONV ela4 THENC - RAND_CONV (RATOR_CONV(RAND_CONV(REWR_CONV labs_def)) THENC eval) THENC - REWR_CONV LET_THM THENC BETA_CONV THENC - RATOR_CONV(RATOR_CONV(RAND_CONV eval)))) th - val tm = th |> rconc |> rator |> rator |> rand - in - if same_const T tm then - CONV_RULE(RAND_CONV (REWR_CONV COND_T THENC add_pos_conv)) th - else (assert (same_const F) tm; - CONV_RULE(RAND_CONV (REWR_CONV COND_F THENC - RAND_CONV eval THENC REWR_CONV LET_THM THENC BETA_CONV THENC - RAND_CONV eval THENC REWR_CONV LET_THM THENC BETA_CONV THENC - add_pos_conv THENC - RAND_CONV(RAND_CONV(RAND_CONV eval THENC REWR_CONV AND_T)))) th) - end - in - f th - end - end -in f end -*) - -(* -fun enc_lines_again_rule labs_def = -let - fun f th = - let - val (th,b) = - (* Asm *) - (CONV_RULE(RAND_CONV (REWR_CONV ela3 THENC add_pos_conv)) th,true) - handle HOL_ERR _ => - (* Label *) - (CONV_RULE(RAND_CONV (REWR_CONV ela2 THENC add_pos_conv)) th,true) - handle HOL_ERR _ => - (* LabAsm *) - let - val th = CONV_RULE(RAND_CONV ( - REWR_CONV ela4 THENC - RAND_CONV (RATOR_CONV(RAND_CONV(REWR_CONV labs_def)) THENC eval) THENC - let_CONV THENC - RATOR_CONV(RATOR_CONV(RAND_CONV eval)))) th - in - (* no offset update *) - (CONV_RULE(RAND_CONV (REWR_CONV COND_T THENC add_pos_conv)) th,true) - handle HOL_ERR _ => - (* offset update *) - (CONV_RULE(RAND_CONV (REWR_CONV COND_F THENC - RAND_CONV eval THENC let_CONV THENC - RAND_CONV eval THENC let_CONV THENC - add_pos_conv THENC - RAND_CONV(RAND_CONV(RAND_CONV eval THENC REWR_CONV AND_T)))) th,true) - end - handle HOL_ERR _ => - (* nil *) - (CONV_RULE(RAND_CONV (REWR_CONV ela1 THENC RATOR_CONV(RAND_CONV eval))) th,false) - in if b then f th else th end -in f end - -fun enc_lines_again_conv labs_def = enc_lines_again_rule labs_def o REFL -*) - -(* -fun enc_lines_again_conv labs_def tm = tm |> ( - IFC - ((REWR_CONV ela3 THENC add_pos_conv) ORELSEC - (REWR_CONV ela2 THENC add_pos_conv) ORELSEC - (REWR_CONV ela4 THENC - RAND_CONV (RATOR_CONV(RAND_CONV(REWR_CONV labs_def)) THENC eval) THENC - let_CONV THENC - RATOR_CONV(RATOR_CONV(RAND_CONV eval)) THENC - ((REWR_CONV COND_T THENC add_pos_conv) ORELSEC - (REWR_CONV COND_F THENC - RAND_CONV eval THENC let_CONV THENC - RAND_CONV eval THENC let_CONV THENC - add_pos_conv THENC - RAND_CONV(RAND_CONV(RAND_CONV eval THENC REWR_CONV AND_T)))))) - (enc_lines_again_conv labs_def) - (REWR_CONV ela1 THENC RATOR_CONV(RAND_CONV eval))) -*) - -(* -fun enc_lines_again_conv labs_def tm = tm |> ( - IFC - ((REWR_CONV ela3) ORELSEC - (REWR_CONV ela2) ORELSEC - (REWR_CONV ela4 THENC - RAND_CONV (RATOR_CONV(RAND_CONV(REWR_CONV labs_def)) THENC eval) THENC - let_CONV THENC - RATOR_CONV(RATOR_CONV(RAND_CONV eval)) THENC - ((REWR_CONV COND_T) ORELSEC - (REWR_CONV COND_F THENC - RAND_CONV eval THENC let_CONV THENC - RAND_CONV eval THENC let_CONV)))) - (enc_lines_again_conv labs_def) - (REWR_CONV ela1 THENC eval)) -*) - -(* -fun enc_lines_again_conv tm = tm |> ( - IFC - ((REWR_CONV ela3) ORELSEC - (REWR_CONV ela2) ORELSEC - (REWR_CONV ela4 THENC - RAND_CONV eval THENC let_CONV THENC - RATOR_CONV(RATOR_CONV(RAND_CONV eval)) THENC - ((REWR_CONV COND_T) ORELSEC - (REWR_CONV COND_F THENC - RAND_CONV eval THENC let_CONV THENC - RAND_CONV eval THENC let_CONV)))) - (enc_lines_again_conv) - (REWR_CONV ela1 THENC eval)) -*) - -(* -fun enc_lines_again_conv tm = tm |> ( - IFC - ((REWR_CONV ela3) ORELSEC - (REWR_CONV ela2) ORELSEC - (REWR_CONV ela4 THENC - RAND_CONV eval THENC let_CONV THENC - RATOR_CONV(RATOR_CONV(RAND_CONV eval)) THENC - COND_CONV THENC - TRY_CONV - (RAND_CONV eval THENC let_CONV THENC - RAND_CONV eval THENC let_CONV))) - (enc_lines_again_conv) - (REWR_CONV ela1 THENC eval)) -*) - -(* -fun enc_lines_again_conv tm = tm |> ( - IFC - ((REWR_CONV ela3) ORELSEC - (REWR_CONV ela2) ORELSEC - (REWR_CONV ela4 THENC - RAND_CONV eval THENC let_CONV THENC - RATOR_CONV(RATOR_CONV(RAND_CONV eval)) THENC - COND_CONV THENC REPEATC let_CONV)) - (enc_lines_again_conv) - (REWR_CONV ela1 THENC eval)) -*) - -(* -val encoded_prog_thm_prefix = - let - val (car,cdr) = dest_comb (concl encoded_prog_thm) - val (ls,ty) = listSyntax.dest_list cdr - val ls' = List.take(ls,10) - in mk_thm([],mk_comb(car,listSyntax.mk_list(ls',ty))) end -val encoded_prog_defs_prefix = - List.take(List.rev encoded_prog_defs,10) - -(* rule using exceptions *) -val enc_secs_again_thm = - tm13 |> timez "enc_secs_again" ( - RAND_CONV(REWR_CONV encoded_prog_thm_prefix) THENC - enc_secs_again_conv - "enc_again_" (enc_lines_again_conv computed_labs_def) 0 - encoded_prog_defs_prefix) -2m15s - -(* custom conv *) -val enc_secs_again_thm = - tm13 |> timez "enc_secs_again" ( - RAND_CONV(REWR_CONV encoded_prog_thm_prefix) THENC - enc_secs_again_conv - "enc_again_" (enc_lines_again_conv computed_labs_def) 0 - encoded_prog_defs_prefix) -2m17s - -(* custom conv with IFC *) -val enc_secs_again_thm = - tm13 |> timez "enc_secs_again" ( - RAND_CONV(REWR_CONV encoded_prog_thm_prefix) THENC - enc_secs_again_conv - "enc_again_" (enc_lines_again_conv computed_labs_def) 0 - encoded_prog_defs_prefix) -2m18s - -val enc_secs_again_thm = - tm13 |> timez "enc_secs_again" ( - RAND_CONV(REWR_CONV encoded_prog_thm_prefix) THENC - enc_secs_again_conv - "enc_again_" - (PATH_CONV "llllr" (REWR_CONV computed_labs_def) THENC eval) - 0 - encoded_prog_defs_prefix) -14.4s - -val (dth::dths) = encoded_prog_defs_prefix -val tm = tm13 |> RAND_CONV(REWR_CONV encoded_prog_thm_prefix) |> rconc -val th1 = tm |> (REWR_CONV esc THENC (RAND_CONV(RATOR_CONV(RAND_CONV(REWR_CONV dth))))) -val tm = th1 |> rconc |> rand - -val test1 = - time (PATH_CONV "llllr" (REWR_CONV computed_labs_def) THENC eval) tm -1.1s - -val th = REFL tm - -(* rule with exceptions *) -val test2 = - time (enc_lines_again_conv computed_labs_def) tm -17.1s -val check = rconc test1 = rconc test2 - -(* custom conv with IFC *) -val test3 = - time (enc_lines_again_conv computed_labs_def) tm -30.6s -val check = rconc test1 = rconc test3 - -(* custom conv with IFC and delayed leaves *) -val test4 = - time (enc_lines_again_conv computed_labs_def) tm -31.3s -val check = rconc test1 = rconc test4 - -(* custom conv with labs expanded *) -val test5 = - time (PATH_CONV "llllr" (REWR_CONV computed_labs_def) THENC enc_lines_again_conv) tm -21.2s -val check = rconc test1 = rconc test5 - -(* custom conv with labs expanded using COND_CONV *) -val test6 = - time (PATH_CONV "llllr" (REWR_CONV computed_labs_def) THENC enc_lines_again_conv) tm -19.4s -val check = rconc test1 = rconc test6 - -(* custom conv with labs expanded using COND_CONV and deferring all computations *) -val test7 = - time (PATH_CONV "llllr" (REWR_CONV computed_labs_def) THENC enc_lines_again_conv) tm -19.6s -val check = rconc test1 = rconc test7 - -*) +val _ = export_theory (); diff --git a/compiler/bootstrap/compilation/x64/64/x64_config_encScript.sml b/compiler/bootstrap/compilation/x64/64/x64_config_encScript.sml deleted file mode 100644 index 9c4bef5bde..0000000000 --- a/compiler/bootstrap/compilation/x64/64/x64_config_encScript.sml +++ /dev/null @@ -1,49 +0,0 @@ -(* - Encoding of compiler state -*) -open preamble backend_enc_decTheory x64BootstrapTheory; - -val _ = new_theory "x64_config_enc"; - -Theorem tap_config_enc'_comp[compute]: - tap_config_enc' c = Tree 0 [bool_enc' (c.explore_flag)] -Proof - Cases_on ‘c’ - \\ fs [backend_enc_decTheory.tap_config_enc'_def] -QED - -val confs = LIST_CONJ - [to_dataBootstrapTheory.flat_conf_def, - to_dataBootstrapTheory.bvl_conf_def, - to_dataBootstrapTheory.bvi_conf_def] - -val encode_backend_config_cake_config_lemma = - “encode_backend_config $ config_to_inc_config $ cake_config with <| - word_to_word_conf := (cake_config.word_to_word_conf with reg_alg := 4) |>” - |> (SIMP_CONV (srw_ss()) [cake_config_def,confs,encode_backend_config_def] THENC EVAL); - -Definition config_enc_str_def: - config_enc_str = ^(encode_backend_config_cake_config_lemma |> concl |> rand) -End - -Theorem encode_backend_config_cake_config = - encode_backend_config_cake_config_lemma - |> CONV_RULE (RAND_CONV (REWR_CONV (GSYM config_enc_str_def))); - -val _ = let - val filename = "config_enc_str.txt" -(* config_enc_str_def |> concl |> rand |> repeat (snd o stringSyntax.dest_string) *) - val str = config_enc_str_def |> concl |> rand |> stringSyntax.fromHOLstring - val n = size str - fun nice_size_str n = - if n < 1024 then int_to_string n ^ " bytes" else - if n < 1024 * 1024 then int_to_string (n div 1024) ^ " kilobytes" else - int_to_string (n div (1024 * 1024)) ^ " megabytes" - val _ = print ("Writing " ^ nice_size_str n ^ " to " ^ filename ^ " .. ") - val f = TextIO.openOut filename - val _ = TextIO.output(f,str) - val _ = TextIO.closeOut f - val _ = print "done.\n" - in () end; - -val _ = export_theory(); diff --git a/compiler/bootstrap/translation/compiler32ProgScript.sml b/compiler/bootstrap/translation/compiler32ProgScript.sml index b25d83d87f..c4adf0be09 100644 --- a/compiler/bootstrap/translation/compiler32ProgScript.sml +++ b/compiler/bootstrap/translation/compiler32ProgScript.sml @@ -257,7 +257,7 @@ val res = translate (extend_conf_def |> spec32 |> SIMP_RULE (srw_ss()) [MEMBER_I val res = translate parse_target_32_def; val res = translate add_tap_output_def; -val res = translate ffinames_to_string_list_def; +val res = translate backendTheory.ffinames_to_string_list_def; val res = format_compiler_result_def |> Q.GENL[`bytes`,`c`] diff --git a/compiler/bootstrap/translation/compiler64ProgScript.sml b/compiler/bootstrap/translation/compiler64ProgScript.sml index 3d098de97b..90b94a3d03 100644 --- a/compiler/bootstrap/translation/compiler64ProgScript.sml +++ b/compiler/bootstrap/translation/compiler64ProgScript.sml @@ -282,7 +282,7 @@ val res = format_compiler_result_def |> spec64 |> translate; -val res = translate ffinames_to_string_list_def; +val res = translate backendTheory.ffinames_to_string_list_def; val res = translate compile_64_def; diff --git a/compiler/compilerScript.sml b/compiler/compilerScript.sml index 7e3288ab41..fdbd6308dd 100644 --- a/compiler/compilerScript.sml +++ b/compiler/compilerScript.sml @@ -634,14 +634,6 @@ val add_tap_output_def = Define` add_tap_output td out = if td = Nil then out else td :mlstring app_list`; -Definition ffinames_to_string_list_def: - (ffinames_to_string_list [] = []) ∧ - (ffinames_to_string_list ((ExtCall s)::rest) = - s::(ffinames_to_string_list rest)) ∧ - (ffinames_to_string_list ((SharedMem _)::rest) = - ffinames_to_string_list rest) -End - (* The top-level compiler with everything instantiated except it doesn't do exporting *) (* The top-level compiler with almost everything instantiated except the top-level configuration *) diff --git a/compiler/inference/Holmakefile b/compiler/inference/Holmakefile index e3f19df0fb..147cde8087 100644 --- a/compiler/inference/Holmakefile +++ b/compiler/inference/Holmakefile @@ -3,7 +3,7 @@ HOLUNIFDIR = $(HOLDIR)/examples/algorithms/unification/triangular/first-order INCLUDES = $(HOLUNIFDIR) $(HOLUNIFDIR)/compilation \ $(CAKEMLDIR)/misc $(CAKEMLDIR)/semantics \ $(CAKEMLDIR)/semantics/proofs $(CAKEMLDIR)/basis/pure \ - ../../cv_translator + ../../translator all: $(DEFAULT_TARGETS) README.md .PHONY: all diff --git a/compiler/inference/infer_cvScript.sml b/compiler/inference/infer_cvScript.sml index d1e9f99c54..95265d6f94 100644 --- a/compiler/inference/infer_cvScript.sml +++ b/compiler/inference/infer_cvScript.sml @@ -3,7 +3,8 @@ *) open preamble miscTheory; open cv_transLib; -open astTheory namespaceTheory inferTheory inferPropsTheory unify_cvTheory; +open astTheory namespaceTheory inferTheory inferPropsTheory; +open basis_cvTheory unify_cvTheory; val _ = new_theory "infer_cv"; @@ -20,8 +21,6 @@ val _ = cv_auto_trans $ expand write_def; val _ = cv_trans $ init_infer_state_def; val _ = cv_trans $ expand init_state_def; -val _ = cv_trans $ expand mlstringTheory.implode_def; -val _ = cv_auto_trans mlstringTheory.concat_def; val _ = cv_auto_trans lookup_st_ex_def; val _ = cv_auto_trans $ expand fresh_uvar_def; @@ -67,27 +66,6 @@ Proof \\ rw [] \\ simp [Once res] QED -val toChar_pre = cv_trans_pre mlintTheory.toChar_def -val num_to_chars_pre = cv_auto_trans_pre mlintTheory.num_to_chars_def; - -Theorem num_to_chars_pre[cv_pre,local]: - ∀a0 a1 a2 a3. num_to_chars_pre a0 a1 a2 a3 -Proof - ho_match_mp_tac mlintTheory.num_to_chars_ind \\ rw [] - \\ rw [] \\ simp [Once num_to_chars_pre] - \\ once_rewrite_tac [toChar_pre] \\ gvs [] \\ rw [] - \\ ‘k MOD 10 < 10’ by gvs [] \\ simp [] -QED - -Triviality Num_ABS: - Num (ABS i) = Num i -Proof - Cases_on ‘i’ \\ gvs [] -QED - -val _ = cv_trans (mlintTheory.toString_def |> SRULE [Num_ABS]); -val _ = cv_trans mlintTheory.num_to_str_def; - val _ = cv_auto_trans infer_tTheory.type_ident_to_string_def; val res = cv_trans_pre infer_tTheory.ty_var_name_def; diff --git a/compiler/inference/proofs/README.md b/compiler/inference/proofs/README.md index 6adefc890f..42c340d223 100644 --- a/compiler/inference/proofs/README.md +++ b/compiler/inference/proofs/README.md @@ -9,10 +9,6 @@ Proves completeness of the type inferencer, i.e. if there is a type for the program, then the type inferencer will find a type (the most general type). -[inferPropsScript.sml](inferPropsScript.sml): -Various lemmas that are handy in the soundness and completeness -proofs of the type inferencer. - [inferSoundScript.sml](inferSoundScript.sml): Proves soundness of the type inferencer: any type assignment produced by the type inferencer is a valid type for the program. diff --git a/compiler/inference/unify_cvScript.sml b/compiler/inference/unify_cvScript.sml index ec9f27c2a0..fbbd04091f 100644 --- a/compiler/inference/unify_cvScript.sml +++ b/compiler/inference/unify_cvScript.sml @@ -3,7 +3,7 @@ *) open preamble; open cv_transLib unifyTheory; -open cv_stdTheory +open cv_stdTheory basis_cvTheory; val _ = new_theory "unify_cv"; diff --git a/cv_translator/Holmakefile b/cv_translator/Holmakefile index 56e3868472..d4b65b6496 100644 --- a/cv_translator/Holmakefile +++ b/cv_translator/Holmakefile @@ -1,5 +1,10 @@ INCLUDES = $(HOLDIR)/examples/bootstrap \ + ../basis/pure \ + ../compiler/inference \ ../compiler/backend \ + ../compiler/backend/x64 \ + ../compiler/backend/arm8 \ + ../compiler/backend/ag32 \ ../compiler/backend/cv_compute all: $(DEFAULT_TARGETS) README.md @@ -9,3 +14,12 @@ README_SOURCES = $(wildcard *Script.sml) $(wildcard *Lib.sml) $(wildcard *Syntax DIRS = $(patsubst bootstrap/,,$(patsubst encoders/,,$(wildcard */))) README.md: $(CAKEMLDIR)/developers/readme_gen readmePrefix $(patsubst %,%readmePrefix,$(DIRS)) $(README_SOURCES) $(protect $(CAKEMLDIR)/developers/readme_gen) $(README_SOURCES) + +EXTRA_CLEANS = cake_compile_heap + +ifdef POLY +all: cake_compile_heap + +cake_compile_heap: eval_cake_compile_x64Lib.uo eval_cake_compile_arm8Lib.uo + $(HOLDIR)/bin/buildheap -o $@ $< +endif diff --git a/cv_translator/README.md b/cv_translator/README.md index 54f6e661ef..b71e35a651 100644 --- a/cv_translator/README.md +++ b/cv_translator/README.md @@ -18,5 +18,21 @@ Translate non-target-specific backend functions to cv equations. [backend_x64_cvScript.sml](backend_x64_cvScript.sml): Translate x64-specialised functions to cv equations. +[backend_x64_evalScript.sml](backend_x64_evalScript.sml): +Experiments with evaluating the compiler using cv_compute + +[eval_cake_compileLib.sml](eval_cake_compileLib.sml): +Functions for in-logic evaluation of the CakeML compiler. +These use HOL's cv_compute facility. + +[eval_cake_compile_ag32Lib.sml](eval_cake_compile_ag32Lib.sml): +Functions for in-logic evaluation of the CakeML compiler for ag32. + +[eval_cake_compile_arm8Lib.sml](eval_cake_compile_arm8Lib.sml): +Functions for in-logic evaluation of the CakeML compiler for arm8. + +[eval_cake_compile_x64Lib.sml](eval_cake_compile_x64Lib.sml): +Functions for in-logic evaluation of the CakeML compiler for x64. + [to_data_cvScript.sml](to_data_cvScript.sml): Translation of the to_data compiler function. diff --git a/cv_translator/backend_ag32_cvScript.sml b/cv_translator/backend_ag32_cvScript.sml index 9ca7349f5a..63d831584d 100644 --- a/cv_translator/backend_ag32_cvScript.sml +++ b/cv_translator/backend_ag32_cvScript.sml @@ -3,6 +3,7 @@ *) open preamble cv_transLib cv_stdTheory backend_cvTheory backend_32_cvTheory; open backend_ag32Theory ag32Theory ag32_targetTheory to_data_cvTheory; +open export_ag32Theory ag32_configTheory; val _ = new_theory "backend_ag32_cv"; @@ -141,10 +142,26 @@ val _ = cv_trans (compile_0_ag32_def val _ = cv_trans backend_ag32Theory.to_word_0_ag32_def; val _ = cv_auto_trans backend_ag32Theory.to_livesets_0_ag32_def; +(* export *) + +val _ = cv_auto_trans + (export_ag32Theory.ag32_export_def + |> REWRITE_RULE [to_words_line_word, + to_words_line_byte, + split16_eq_chunks16]); + (* main two translations below *) val _ = cv_trans backend_ag32Theory.to_livesets_ag32_def; val _ = cv_trans backend_ag32Theory.compile_cake_ag32_def; +(* lemma used by automation *) + +Theorem set_asm_conf_ag32_backend_config: + set_asm_conf ag32_backend_config ag32_config = ag32_backend_config +Proof + irule backendTheory.set_asm_conf_id \\ EVAL_TAC +QED + val _ = Feedback.set_trace "TheoryPP.include_docs" 0; val _ = export_theory(); diff --git a/cv_translator/backend_arm8_cvScript.sml b/cv_translator/backend_arm8_cvScript.sml index c3b4f6d03b..3bc135fe1a 100644 --- a/cv_translator/backend_arm8_cvScript.sml +++ b/cv_translator/backend_arm8_cvScript.sml @@ -3,6 +3,7 @@ *) open preamble cv_transLib cv_stdTheory backend_cvTheory backend_64_cvTheory; open backend_arm8Theory arm8Theory arm8_targetTheory to_data_cvTheory; +open export_arm8Theory arm8_configTheory; val _ = new_theory "backend_arm8_cv"; @@ -250,10 +251,26 @@ val _ = cv_trans (compile_0_arm8_def val _ = cv_trans backend_arm8Theory.to_word_0_arm8_def; val _ = cv_auto_trans backend_arm8Theory.to_livesets_0_arm8_def; +(* export *) + +val _ = cv_auto_trans + (export_arm8Theory.arm8_export_def + |> REWRITE_RULE [to_words_line_word, + to_words_line_byte, + split16_eq_chunks16]); + (* main two translations below *) val _ = cv_trans backend_arm8Theory.to_livesets_arm8_def; val _ = cv_trans backend_arm8Theory.compile_cake_arm8_def; +(* lemma used by automation *) + +Theorem set_asm_conf_arm8_backend_config: + set_asm_conf arm8_backend_config arm8_config = arm8_backend_config +Proof + irule backendTheory.set_asm_conf_id \\ EVAL_TAC +QED + val _ = Feedback.set_trace "TheoryPP.include_docs" 0; val _ = export_theory(); diff --git a/cv_translator/backend_cvScript.sml b/cv_translator/backend_cvScript.sml index 76b91b4cfd..d1652f0a2e 100644 --- a/cv_translator/backend_cvScript.sml +++ b/cv_translator/backend_cvScript.sml @@ -2,7 +2,8 @@ Translate non-target-specific backend functions to cv equations. *) open preamble cv_transLib cv_stdTheory; -open backendTheory to_data_cvTheory; +open backendTheory to_data_cvTheory exportTheory; +open unify_cvTheory infer_cvTheory basis_cvTheory; val _ = new_theory "backend_cv"; @@ -727,5 +728,141 @@ QED val _ = word_allocTheory.canonize_moves_def |> SRULE [GSYM QSORT_canonize_def] |> cv_auto_trans; +val _ = cv_trans backendTheory.inc_set_oracle_def; + +val _ = cv_trans (exportTheory.escape_sym_char_def |> SRULE [GREATER_EQ]); +val _ = cv_auto_trans exportTheory.emit_symbol_def; +val _ = cv_auto_trans exportTheory.emit_symbols_def; +val _ = cv_trans (exportTheory.data_section_def |> SRULE []); +val _ = cv_trans (exportTheory.data_buffer_def |> SRULE []); +val _ = cv_trans (exportTheory.code_buffer_def |> SRULE []); + +Triviality eq_toChar: + ∀n. n < 16 ⇒ EL n "0123456789ABCDEF" = toChar n +Proof + Cases \\ gvs [] \\ EVAL_TAC + \\ ntac 10 (Cases_on ‘n'’ \\ gvs [] \\ Cases_on ‘n’ \\ gvs []) +QED + +Theorem byte_to_string_eq_toChar_toChar: + byte_to_string b = + strlit (STRING #"0" (STRING #"x" (STRING (toChar (w2n b DIV 16)) + (STRING (toChar (w2n b MOD 16)) [])))) +Proof + simp [exportTheory.byte_to_string_def] + \\ DEP_REWRITE_TAC [eq_toChar] + \\ gvs [DIV_LT_X] + \\ irule LESS_LESS_EQ_TRANS + \\ irule_at Any w2n_lt \\ gvs [] +QED + +val pre = cv_trans_pre byte_to_string_eq_toChar_toChar; +Theorem export_byte_to_string_pre[cv_pre]: + ∀b. export_byte_to_string_pre b +Proof + simp [pre] \\ gen_tac + \\ rpt $ irule_at Any basis_cvTheory.IMP_toChar_pre \\ gvs [] + \\ gvs [DIV_LT_X] + \\ irule LESS_LESS_EQ_TRANS + \\ irule_at Any w2n_lt \\ gvs [] +QED + +val _ = cv_trans exportTheory.preamble_def + +Definition chunks16_def: + chunks16 f xs = + case xs of + [] => Nil + | (x0::x1::x2::x3::x4::x5::x6::x7:: + x8::x9::x10::x11::x12::x13::x14::x15::ys) => + SmartAppend + (f [x0;x1;x2;x3;x4;x5;x6;x7;x8;x9;x10;x11;x12;x13;x14;x15]) + (chunks16 f ys) + | other => f other +End + +Theorem split16_eq_chunks16: + ∀f xs. split16 f xs = chunks16 f xs +Proof + rpt gen_tac + \\ completeInduct_on ‘LENGTH xs’ + \\ rw [] \\ gvs [PULL_FORALL] + \\ simp [Once chunks16_def] + \\ every_case_tac \\ gvs [] + \\ simp [Once split16_def] + \\ simp [Once chunks16_def,SmartAppend_Nil] +QED + +Definition comma_cat_byte_to_string_def: + comma_cat_byte_to_string x = + case x of + [] => [newl_strlit] + | [x] => [byte_to_string x; newl_strlit] + | (x::xs) => byte_to_string x::comm_strlit::comma_cat_byte_to_string xs +End + +Definition comma_cat_word_to_string_def: + comma_cat_word_to_string x = + case x of + [] => [newl_strlit] + | [x] => [word_to_string x; newl_strlit] + | (x::xs) => word_to_string x::comm_strlit::comma_cat_word_to_string xs +End + +val _ = cv_trans word_to_string_def; +val _ = cv_trans newl_strlit_def; +val _ = cv_trans comm_strlit_def; +val _ = cv_trans comma_cat_byte_to_string_def; +val _ = cv_trans comma_cat_word_to_string_def; + +Theorem to_comma_cat_byte_to_string: + ∀xs. comma_cat byte_to_string xs = + comma_cat_byte_to_string xs +Proof + ho_match_mp_tac comma_cat_byte_to_string_ind + \\ rw [] + \\ simp [Once comma_cat_def, Once comma_cat_byte_to_string_def] + \\ Cases_on ‘xs’ \\ gvs [] + \\ Cases_on ‘t’ \\ gvs [] +QED + +Theorem to_comma_cat_word_to_string: + ∀xs. comma_cat word_to_string xs = + comma_cat_word_to_string xs +Proof + ho_match_mp_tac comma_cat_word_to_string_ind + \\ rw [] + \\ simp [Once comma_cat_def, Once comma_cat_word_to_string_def] + \\ Cases_on ‘xs’ \\ gvs [] + \\ Cases_on ‘t’ \\ gvs [] +QED + +Definition words_line_byte_def: + words_line_byte word_directive ls = + List (word_directive :: comma_cat_byte_to_string ls) +End + +Definition words_line_word_def: + words_line_word word_directive ls = + List (word_directive :: comma_cat_word_to_string ls) +End + +val _ = cv_trans words_line_byte_def; +val _ = cv_trans words_line_word_def; + +Theorem to_words_line_byte: + words_line wd byte_to_string = words_line_byte wd +Proof + gvs [FUN_EQ_THM,words_line_byte_def, + to_comma_cat_byte_to_string,words_line_def] +QED + +Theorem to_words_line_word: + words_line wd word_to_string = words_line_word wd +Proof + gvs [FUN_EQ_THM,words_line_word_def, + to_comma_cat_word_to_string,words_line_def] +QED + val _ = Feedback.set_trace "TheoryPP.include_docs" 0; val _ = export_theory(); diff --git a/cv_translator/backend_x64_cvScript.sml b/cv_translator/backend_x64_cvScript.sml index 497871e17c..5af3638019 100644 --- a/cv_translator/backend_x64_cvScript.sml +++ b/cv_translator/backend_x64_cvScript.sml @@ -3,6 +3,7 @@ *) open preamble cv_transLib cv_stdTheory backend_cvTheory backend_64_cvTheory; open backend_x64Theory x64Theory x64_targetTheory to_data_cvTheory; +open export_x64Theory x64_configTheory; val _ = new_theory "backend_x64_cv"; @@ -236,10 +237,26 @@ val _ = cv_trans (compile_0_x64_def val _ = cv_trans backend_x64Theory.to_word_0_x64_def; val _ = cv_auto_trans backend_x64Theory.to_livesets_0_x64_def; +(* export *) + +val _ = cv_auto_trans + (export_x64Theory.x64_export_def + |> REWRITE_RULE [to_words_line_word, + to_words_line_byte, + split16_eq_chunks16]); + (* main two translations below *) val _ = cv_trans backend_x64Theory.to_livesets_x64_def; val _ = cv_trans backend_x64Theory.compile_cake_x64_def; +(* lemma used by automation *) + +Theorem set_asm_conf_x64_backend_config: + set_asm_conf x64_backend_config x64_config = x64_backend_config +Proof + irule backendTheory.set_asm_conf_id \\ EVAL_TAC +QED + val _ = Feedback.set_trace "TheoryPP.include_docs" 0; val _ = export_theory(); diff --git a/cv_translator/backend_x64_evalScript.sml b/cv_translator/backend_x64_evalScript.sml new file mode 100644 index 0000000000..f4232308d3 --- /dev/null +++ b/cv_translator/backend_x64_evalScript.sml @@ -0,0 +1,16 @@ +(* + Experiments with evaluating the compiler using cv_compute +*) +open preamble; + +val _ = new_theory "backend_x64_eval"; + +open eval_cake_compile_x64Lib; + +Definition prog_def: + prog = [] : ast$dec list +End + +val res = eval_cake_compile_x64 "" prog_def "empty_prog.S"; + +val _ = export_theory(); diff --git a/cv_translator/eval_cake_compileLib.sig b/cv_translator/eval_cake_compileLib.sig new file mode 100644 index 0000000000..25743a0976 --- /dev/null +++ b/cv_translator/eval_cake_compileLib.sig @@ -0,0 +1,31 @@ + +signature eval_cake_compileLib = +sig + + include Abbrev + + type arch_thms = + { default_config_def : thm + , default_config_simp : thm + , to_livesets_def : thm + , compile_cake_def : thm + , compile_cake_imp : thm + , cv_export_def : thm } + + type comp_input = + { prefix : string + , conf_def : thm + , prog_def : thm + , output_filename : string + , output_conf_filename : string option } + + val eval_cake_compile : + arch_thms -> string -> thm -> string -> thm + + val eval_cake_compile_with_conf : + arch_thms -> string -> thm -> thm -> string -> thm + + val eval_cake_compile_general : + arch_thms -> comp_input -> thm + +end diff --git a/cv_translator/eval_cake_compileLib.sml b/cv_translator/eval_cake_compileLib.sml new file mode 100644 index 0000000000..1fdd42a803 --- /dev/null +++ b/cv_translator/eval_cake_compileLib.sml @@ -0,0 +1,159 @@ +(* + Functions for in-logic evaluation of the CakeML compiler. + These use HOL's cv_compute facility. +*) +structure eval_cake_compileLib :> eval_cake_compileLib = +struct + +open preamble cv_transLib; +open exportTheory; +open to_data_cvTheory; + +type arch_thms = + { default_config_def : thm + , default_config_simp : thm + , to_livesets_def : thm + , compile_cake_def : thm + , compile_cake_imp : thm + , cv_export_def : thm } + +type comp_input = + { prefix : string + , conf_def : thm + , prog_def : thm + , output_filename : string + , output_conf_filename : string option } + +fun write_cv_char_list_to_file filename cv_char_list_tm = let + val f = TextIO.openOut filename + fun loop tm = let + val (n,rest) = cvSyntax.dest_cv_pair tm + val c = cvSyntax.dest_cv_num n |> numSyntax.int_of_term |> chr + val _ = TextIO.output1(f,c) + in loop rest end handle HOL_ERR _ => (); + val _ = loop cv_char_list_tm + in TextIO.closeOut f end; + +fun allowing_rebind f = Feedback.trace ("Theory.allow_rebinds", 1) f; + +fun eval_cake_compile_general (arch : arch_thms) (input : comp_input) = let + val { prefix, conf_def, prog_def + , output_filename , output_conf_filename } = input + val { default_config_def, default_config_simp, to_livesets_def + , compile_cake_def, compile_cake_imp, cv_export_def } = arch + fun define_abbrev name tm = + Feedback.trace ("Theory.allow_rebinds", 1) + (mk_abbrev (prefix ^ name)) tm + val conf = conf_def |> concl |> lhs + val c = backendTheory.config_to_inc_config_def + |> ISPEC conf |> CONV_RULE (RAND_CONV EVAL) + val _ = allowing_rebind (cv_trans_deep_embedding EVAL) prog_def + val input_tm = to_livesets_def |> GEN_ALL + |> SPEC (prog_def |> concl |> lhs) + |> SPEC (c |> concl |> rand) |> concl |> lhs |> mk_fst + val oracles = let + val graphs = cv_eval input_tm |> rconc + in reg_allocComputeLib.get_oracle reg_alloc.Irc graphs end + val oracle_def = define_abbrev "temp_oracle" oracles; + val _ = allowing_rebind (cv_trans_deep_embedding EVAL) oracle_def + val oracle_tm = oracle_def |> concl |> lhs + val c_tm = c |> concl |> lhs + val c_oracle_tm = backendTheory.inc_set_oracle_def + |> SPEC (c |> concl |> rhs) + |> SPEC oracle_tm |> concl |> lhs + val input_tm = compile_cake_def |> GEN_ALL + |> SPEC (prog_def |> concl |> lhs) + |> SPEC c_oracle_tm |> concl |> lhs + val to_option_some = cv_typeTheory.to_option_def |> cj 2 + val to_pair = cv_typeTheory.to_pair_def |> cj 1 + val th1 = cv_eval_raw input_tm + |> CONV_RULE (PATH_CONV "lr" (REWRITE_CONV [GSYM c])) + val th2 = th1 |> CONV_RULE (PATH_CONV "r" (REWR_CONV to_option_some)) + handle HOL_ERR _ => failwith "compiler returned NONE" + val c2n_Num = cvTheory.c2n_def |> cj 1 + val th3 = th2 |> CONV_RULE (PATH_CONV "rr" (REWR_CONV to_pair) + THENC PATH_CONV "rrr" (REWR_CONV to_pair) + THENC PATH_CONV "rrrr" (REWR_CONV to_pair) + THENC PATH_CONV "rrrrr" (REWR_CONV to_pair) + THENC PATH_CONV "rrrrrr" (REWR_CONV to_pair) + THENC PATH_CONV "rrrrrrr" (REWR_CONV to_pair) + THENC PATH_CONV "rrrrrrrr" (REWR_CONV to_pair) + THENC PATH_CONV "rrrlr" (REWR_CONV c2n_Num) + THENC PATH_CONV "rrrrrlr" (REWR_CONV c2n_Num) + THENC PATH_CONV "rrrrrrrlr" (REWR_CONV c2n_Num)) + fun abbrev_inside name path th = let + val tm = dest_path path (concl th) + val def = define_abbrev name tm + in (def, CONV_RULE (PATH_CONV path (REWR_CONV (SYM def))) th) end + val (code_def,th) = abbrev_inside "code" "rrlr" th3 + val (data_def,th) = abbrev_inside "data" "rrrrlr" th + val (ffis_def,th) = abbrev_inside "ffis" "rrrrrrlr" th + val (syms_def,th) = abbrev_inside "syms" "rrrrrrrrlr" th + val (conf_def,th) = abbrev_inside "conf" "rrrrrrrrr" th + fun new_spec th = + new_specification(prefix ^ "compiled", + [prefix ^ "oracle", prefix ^ "info"], th) + val result_th = MATCH_MP compile_cake_imp th + |> REWRITE_RULE [backendTheory.inc_set_oracle_pull, + backendTheory.inc_config_to_config_config_to_inc_config] + |> REWRITE_RULE [default_config_simp,LENGTH_NIL] + |> CONV_RULE (UNBETA_CONV oracle_tm) + |> MATCH_MP backend_asmTheory.exists_oracle + |> CONV_RULE (PATH_CONV "b" BETA_CONV) + |> new_spec + (* --- *) + val e = cv_export_def |> concl |> strip_forall |> snd |> lhs + val cv_ty = cvSyntax.cv + fun get_one_subst name abbrev_def = + mk_var(name,cvSyntax.cv) |-> (abbrev_def |> concl |> rhs |> rand) + fun cv_explode e = SPEC e basis_cvTheory.cv_mlstring_explode_def |> concl |> lhs + fun cv_concat e = SPEC e basis_cvTheory.cv_mlstring_concat_def |> concl |> lhs + fun cv_append e = SPEC e basis_cvTheory.cv_misc_append_def |> concl |> lhs + val export_tm = e |> cv_append |> cv_concat |> cv_explode |> subst + [get_one_subst "cv_ffi_names" ffis_def, + get_one_subst "cv_data" data_def, + get_one_subst "cv_bytes" code_def, + get_one_subst "cv_syms" syms_def] + val _ = null (free_vars export_tm) orelse failwith "failed to eval export" + (* + cv_repLib.cv_rep_for [] + “explode $ concat $ misc$append (x64_export x64_ffis x64_code x64_data x64_syms)” + *) + val l = cv_computeLib.cv_compute (cv_transLib.cv_eqs_for export_tm) export_tm + |> concl |> rhs + val _ = write_cv_char_list_to_file output_filename l + val _ = case output_conf_filename of NONE => () + | SOME fname => write_cv_char_list_to_file fname + (conf_def |> concl |> rhs |> rand) + (* --- *) + val _ = Theory.delete_binding (prefix ^ "temp_oracle_cv_eq") + val _ = Theory.delete_binding (prefix ^ "temp_oracle_cv_def") + val _ = Theory.delete_binding (prefix ^ "temp_oracle_def") + val _ = Theory.delete_binding (prefix ^ "syms_def") + in result_th end + +fun eval_cake_compile_with_conf arch prefix conf_def prog_def filename = + eval_cake_compile_general arch + { prefix = prefix + , conf_def = conf_def + , prog_def = prog_def + , output_filename = filename + , output_conf_filename = NONE }; + +fun eval_cake_compile arch prefix = + eval_cake_compile_with_conf arch prefix (#default_config_def arch); + +val _ = Feedback.set_trace "TheoryPP.include_docs" 0; + +(* --- for debugging --- + val _ = (max_print_depth := 15); + val arch = x64_arch_thms; + val input = + { prefix = "x64_" + , conf_def = #default_config_def x64_arch_thms + , prog_def = Define `prog = [] : ast$dec list` + , output_filename = "test.S" + , output_conf_filename = SOME "test_conf.txt" } : comp_input; +*) + +end diff --git a/cv_translator/eval_cake_compile_ag32Lib.sml b/cv_translator/eval_cake_compile_ag32Lib.sml new file mode 100644 index 0000000000..404f0315a5 --- /dev/null +++ b/cv_translator/eval_cake_compile_ag32Lib.sml @@ -0,0 +1,27 @@ +(* + Functions for in-logic evaluation of the CakeML compiler for ag32. +*) +structure eval_cake_compile_ag32Lib = +struct + +local + + open eval_cake_compileLib; + + val ag32_arch_thms = + { default_config_def = ag32_configTheory.ag32_backend_config_def + , default_config_simp = backend_ag32_cvTheory.set_asm_conf_ag32_backend_config + , to_livesets_def = backend_ag32Theory.to_livesets_ag32_def + , compile_cake_def = backend_ag32Theory.compile_cake_ag32_def + , compile_cake_imp = backend_ag32Theory.compile_cake_ag32_thm + , cv_export_def = backend_ag32_cvTheory.cv_ag32_export_def } : arch_thms; + +in + + val eval_cake_compile_ag32 = eval_cake_compile ag32_arch_thms; + val eval_cake_compile_ag32_with_conf = eval_cake_compile_with_conf ag32_arch_thms; + val eval_cake_compile_ag32_general = eval_cake_compile_general ag32_arch_thms; + +end + +end diff --git a/cv_translator/eval_cake_compile_arm8Lib.sml b/cv_translator/eval_cake_compile_arm8Lib.sml new file mode 100644 index 0000000000..33a9f7ae0c --- /dev/null +++ b/cv_translator/eval_cake_compile_arm8Lib.sml @@ -0,0 +1,27 @@ +(* + Functions for in-logic evaluation of the CakeML compiler for arm8. +*) +structure eval_cake_compile_arm8Lib = +struct + +local + + open eval_cake_compileLib; + + val arm8_arch_thms = + { default_config_def = arm8_configTheory.arm8_backend_config_def + , default_config_simp = backend_arm8_cvTheory.set_asm_conf_arm8_backend_config + , to_livesets_def = backend_arm8Theory.to_livesets_arm8_def + , compile_cake_def = backend_arm8Theory.compile_cake_arm8_def + , compile_cake_imp = backend_arm8Theory.compile_cake_arm8_thm + , cv_export_def = backend_arm8_cvTheory.cv_arm8_export_def } : arch_thms; + +in + + val eval_cake_compile_arm8 = eval_cake_compile arm8_arch_thms; + val eval_cake_compile_arm8_with_conf = eval_cake_compile_with_conf arm8_arch_thms; + val eval_cake_compile_arm8_general = eval_cake_compile_general arm8_arch_thms; + +end + +end diff --git a/cv_translator/eval_cake_compile_x64Lib.sml b/cv_translator/eval_cake_compile_x64Lib.sml new file mode 100644 index 0000000000..02374c3172 --- /dev/null +++ b/cv_translator/eval_cake_compile_x64Lib.sml @@ -0,0 +1,27 @@ +(* + Functions for in-logic evaluation of the CakeML compiler for x64. +*) +structure eval_cake_compile_x64Lib = +struct + +local + + open eval_cake_compileLib; + + val x64_arch_thms = + { default_config_def = x64_configTheory.x64_backend_config_def + , default_config_simp = backend_x64_cvTheory.set_asm_conf_x64_backend_config + , to_livesets_def = backend_x64Theory.to_livesets_x64_def + , compile_cake_def = backend_x64Theory.compile_cake_x64_def + , compile_cake_imp = backend_x64Theory.compile_cake_x64_thm + , cv_export_def = backend_x64_cvTheory.cv_x64_export_def } : arch_thms; + +in + + val eval_cake_compile_x64 = eval_cake_compile x64_arch_thms; + val eval_cake_compile_x64_with_conf = eval_cake_compile_with_conf x64_arch_thms; + val eval_cake_compile_x64_general = eval_cake_compile_general x64_arch_thms; + +end + +end diff --git a/cv_translator/to_data_cvScript.sml b/cv_translator/to_data_cvScript.sml index 73f103a5be..0a7efe38ad 100644 --- a/cv_translator/to_data_cvScript.sml +++ b/cv_translator/to_data_cvScript.sml @@ -1,28 +1,2743 @@ (* Translation of the to_data compiler function. *) -open preamble cv_transLib cv_stdTheory; +open preamble cv_transLib cv_stdTheory basis_cvTheory; open backendTheory backend_asmTheory; +open unify_cvTheory infer_cvTheory; val _ = new_theory "to_data_cv"; val _ = cv_memLib.use_long_names := true; -val _ = cv_trans sptreeTheory.fromAList_def; -val _ = cv_trans miscTheory.append_aux_def; -val _ = cv_trans miscTheory.append_def; -val _ = cv_trans miscTheory.tlookup_def; -val _ = cv_trans mlstringTheory.explode_thm; -val _ = cv_trans miscTheory.list_max_def; -val _ = cv_trans (miscTheory.max3_def |> PURE_REWRITE_RULE [GREATER_DEF]); +Triviality list_mem[cv_inline] = listTheory.MEM; -Theorem to_data_fake: - backend_asm$to_data c p = (c,[],LN) +(* source_let *) + +val _ = cv_trans source_letTheory.dest_Letrec_def; +val _ = cv_trans source_letTheory.dest_Let_def; +val _ = cv_trans source_letTheory.lift_let_def; +val _ = cv_trans source_letTheory.lift_lets_def; +val _ = cv_trans source_letTheory.compile_decs_def; + +(* source_to_source *) + +val _ = cv_trans source_to_sourceTheory.compile_def; + +(* source_to_flat *) +val _ = cv_auto_trans namespaceTheory.nsEmpty_def; +val _ = cv_auto_trans namespaceTheory.nsLookup_def; +val _ = cv_auto_trans source_to_flatTheory.empty_env_def; +val _ = cv_trans flatLangTheory.pat_bindings_def; +val _ = cv_trans astTheory.pat_bindings_def; +val _ = cv_trans OPTION_JOIN_DEF; +val _ = cv_trans source_to_flatTheory.type_group_id_type_def; + +Definition compile_pat_alt_def: + (compile_pat_alt env (ast$Pvar v) = flatLang$Pvar v) ∧ + (compile_pat_alt env Pany = Pany) ∧ + (compile_pat_alt env (Plit l) = Plit l) ∧ + (compile_pat_alt env (ast$Pcon id ps) = + flatLang$Pcon + (OPTION_JOIN (OPTION_MAP (nsLookup env.c) id)) + (compile_pats_alt env ps)) ∧ + (compile_pat_alt env (Pref p) = Pref (compile_pat_alt env p)) ∧ + (compile_pat_alt env (Pas p i) = Pas (compile_pat_alt env p) i) ∧ + (compile_pat_alt env (Ptannot p t) = compile_pat_alt env p) ∧ + (compile_pats_alt env [] = []) ∧ + (compile_pats_alt env (p::ps) = compile_pat_alt env p::compile_pats_alt env ps) +End + +val _ = cv_auto_trans (compile_pat_alt_def |> PURE_REWRITE_RULE [oneline OPTION_MAP_DEF,o_THM]); + +Theorem compile_pat_alt_thm: + (∀v env. compile_pat_alt env v = compile_pat env v) ∧ + (∀v env. compile_pats_alt env v = MAP (compile_pat env) v) +Proof + Induct >> rw[compile_pat_alt_def,source_to_flatTheory.compile_pat_def] >> metis_tac[] +QED + +val _ = cv_trans $ GSYM $ cj 1 compile_pat_alt_thm; + +val pre = cv_auto_trans_pre + (source_to_flatTheory.compile_exp_def + |> PURE_REWRITE_RULE [oneline OPTION_MAP_DEF,o_THM]); + +Theorem source_to_flat_compile_exp_pre[cv_pre,local]: + (∀t env v. source_to_flat_compile_exp_pre t env v) ∧ + (∀t env v. source_to_flat_compile_exps_pre t env v) ∧ + (∀t env v. source_to_flat_compile_pes_pre t env v) ∧ + (∀t env v. source_to_flat_compile_funs_pre t env v) +Proof + ho_match_mp_tac source_to_flatTheory.compile_exp_ind >> + rw[] >> + rw[Once pre] >> + PURE_TOP_CASE_TAC >> gvs[source_to_flatTheory.environment_fn_updates] +QED + +Definition make_varls_alt_def: + (make_varls_alt idx [] = Con None NONE []) ∧ + (make_varls_alt idx [x] = App None (GlobalVarInit idx) [Var_local None x]) ∧ + (make_varls_alt idx (x::xs) = + Let None NONE (App None (GlobalVarInit idx) [Var_local None x]) + (make_varls_alt (idx + 1) xs):flatLang$exp) +End + +val _ = cv_trans make_varls_alt_def + +Theorem make_varls_alt_thm: + ∀n tra idx xs. make_varls n tra idx xs = make_varls_alt idx xs +Proof + ho_match_mp_tac source_to_flatTheory.make_varls_ind >> + rw[source_to_flatTheory.make_varls_def,make_varls_alt_def] +QED + +val _ = cv_trans make_varls_alt_thm + +Theorem list_size_mono: + (∀x. m x ≤ m' x) ⇒ list_size m xs ≤ list_size m' xs +Proof + strip_tac >> Induct_on ‘xs’ >> rw[list_size_def] >> + irule LESS_EQ_LESS_EQ_MONO >> + rw[] +QED + +Definition nsMap_alt_def: + nsMap_alt data (Bind v m) = Bind (MAP (λ(n,x). (n,(x,data))) v) (nsMap_alts data m) ∧ + nsMap_alts data [] = [] ∧ + nsMap_alts data ((x,y)::xs) = (x,nsMap_alt data y)::nsMap_alts data xs +Termination + wf_rel_tac ‘measure $ λx. + case x of + INL (f,y) => namespace_size (K 0) (K 0) (K 0) y + | INR (f,xs) => list_size (namespace_size (K 0) (K 0) (K 0) o SND) xs’ >> + rw[namespaceTheory.namespace_size_eq] >> + qmatch_goalsub_abbrev_tac ‘a1 < a2 + _’ >> + ‘a1 ≤ a2’ suffices_by rw[] >> + unabbrev_all_tac >> + irule list_size_mono >> + Cases >> + rw[basicSizeTheory.pair_size_def] +End + +Theorem nsMap_alt_thm: + (∀data b. nsMap_alt (data:'d) (b:('a,'b,'c) namespace) = nsMap (λtag. (tag, data)) b) ∧ + (∀data xs. nsMap_alts (data:'d) (xs:('a # ('a,'b,'c) namespace) list) = (MAP (λ(mn,e). (mn,nsMap (λtag. (tag, data)) e)) $ xs)) +Proof + ho_match_mp_tac nsMap_alt_ind >> + rw[nsMap_alt_def,namespaceTheory.nsMap_def] +QED + +val _ = cv_auto_trans nsMap_alt_def; + +val _ = cv_auto_trans (source_to_flatTheory.alloc_tags_def |> PURE_ONCE_REWRITE_RULE[GSYM nsMap_alt_thm]) + +Definition compile_decs_alt_def: + (compile_dec_alt (t:string list) n next env envs (ast$Dlet locs p e) = + let n' = n + 4 in + let xs = REVERSE (pat_bindings p []) in + let e' = compile_exp (xs++t) env e in + let l = LENGTH xs in + let n'' = n' + l in + (n'', (next with vidx := next.vidx + l), + <| v := alist_to_ns (alloc_defs n' next.vidx xs); c := nsEmpty |>, + envs, + [flatLang$Dlet (Mat None e' + [(compile_pat env p, make_varls 0 None next.vidx xs)])])) ∧ + (compile_dec_alt t n next env envs (ast$Dletrec locs funs) = + let fun_names = MAP FST funs in + let new_env = nsBindList (MAP (\x. (x, Local None x)) fun_names) env.v in + let flat_funs = compile_funs t (env with v := new_env) funs in + let n' = n + 1 in + let env' = <| v := alist_to_ns (alloc_defs n' next.vidx (REVERSE fun_names)); + c := nsEmpty |> in + (n' + LENGTH funs, (next with vidx := next.vidx + LENGTH funs), + env', envs, + [flatLang$Dlet (flatLang$Letrec (join_all_names t) flat_funs + (make_varls 0 None next.vidx (REVERSE fun_names)))])) /\ + (compile_dec_alt t n next env envs (Dtype locs type_def) = + let new_env = MAPi (\tid (_,_,constrs). alloc_tags (next.tidx + tid) constrs) type_def in + (n, (next with tidx := next.tidx + LENGTH type_def), + <| v := nsEmpty; + c := FOLDL (\ns (l,cids). nsAppend l ns) nsEmpty new_env |>, + envs, + MAPi (λi (ns,cids). flatLang$Dtype (next.tidx + i) cids) new_env)) ∧ + (compile_dec_alt _ n next env envs (Dtabbrev locs tvs tn t) = + (n, next, empty_env, envs, [])) ∧ + (compile_dec_alt t n next env envs (Dexn locs cn ts) = + (n, (next with eidx := next.eidx + 1), + <| v := nsEmpty; c := nsSing cn (next.eidx, NONE) |>, + envs, + [Dexn next.eidx (LENGTH ts)])) ∧ + (compile_dec_alt t n next env envs (Dmod mn ds) = + let (n', next', new_env, envs', ds') = compile_decs_alt (mn::t) n next env envs ds in + (n', next', (lift_env mn new_env), envs', ds')) ∧ + (compile_dec_alt t n next env envs (Dlocal lds ds) = + let (n', next1, new_env1, envs', lds') = compile_decs_alt t n next env envs lds in + let (n'', next2, new_env2, envs'', ds') = compile_decs_alt t n' next1 + (extend_env new_env1 env) envs' ds + in (n'', next2, new_env2, envs'', lds'++ds')) ∧ + (compile_dec_alt t n next env envs (Denv nenv) = + (n + 1, next with vidx := next.vidx + 1, + <| v := nsBind nenv (Glob None next.vidx) nsEmpty; c := nsEmpty |>, + envs with <| next := envs.next + 1; + envs := insert envs.next env envs.envs |>, + [flatLang$Dlet (App None (GlobalVarInit next.vidx) + [env_id_tuple envs.generation envs.next])])) ∧ + (compile_decs_alt t n next env envs [] = + (n, next, empty_env, envs, [])) ∧ + (compile_decs_alt t n next env envs (d::ds) = + let (n', next1, new_env1, envs1, d') = compile_dec_alt t n next env envs d in + let (n'', next2, new_env2, envs2, ds') = + compile_decs_alt t n' next1 (extend_env new_env1 env) envs1 ds + in + (n'', next2, extend_env new_env2 new_env1, envs2, d'++ds')) +End + +Theorem compile_decs_cons: + compile_decs t n next env envs (d::ds) = + let (n', next1, new_env1, envs1, d') = compile_decs t n next env envs [d] in + let (n'', next2, new_env2, envs2, ds') = + compile_decs t n' next1 (extend_env new_env1 env) envs1 ds + in + (n'', next2, extend_env new_env2 new_env1, envs2, d'++ds') +Proof + rw[Once $ oneline source_to_flatTheory.compile_decs_def] >> + rpt(PURE_TOP_CASE_TAC >> gvs[source_to_flatTheory.compile_decs_def]) >> + gvs[source_to_flatTheory.extend_env_def,source_to_flatTheory.empty_env_def, + UNCURRY_eq_pair,PULL_EXISTS] >> + qmatch_goalsub_abbrev_tac ‘$= a1’ >> PairCases_on ‘a1’ >> + pop_assum $ assume_tac o GSYM >> + gvs[markerTheory.Abbrev_def] >> + rw[source_to_flatTheory.environment_component_equality] >> + qmatch_goalsub_abbrev_tac ‘$= a1’ >> PairCases_on ‘a1’ >> + pop_assum $ assume_tac o GSYM >> + gvs[markerTheory.Abbrev_def] >> + rw[source_to_flatTheory.environment_component_equality] +QED + +Theorem compile_decs_thm: + ∀t n next env envs xs. + compile_decs t n next env envs xs = compile_decs_alt t n next env envs xs +Proof + ho_match_mp_tac source_to_flatTheory.compile_decs_ind >> + PURE_ONCE_REWRITE_TAC[compile_decs_cons] >> + rw[source_to_flatTheory.compile_decs_def,compile_decs_alt_def,source_to_flatTheory.extend_env_def,source_to_flatTheory.empty_env_def,UNCURRY_eq_pair,PULL_EXISTS,source_to_flatTheory.lift_env_def] + >- metis_tac[FST,SND,PAIR] + >- metis_tac[FST,SND,PAIR] >> + res_tac >> + simp[compile_decs_cons,UNCURRY_eq_pair,PULL_EXISTS,source_to_flatTheory.extend_env_def] +QED + +val _ = cv_auto_trans compile_decs_alt_def + +val _ = cv_trans compile_decs_thm + +val _ = cv_auto_trans source_to_flatTheory.compile_prog_def + +(* flat_pattern *) + +Definition compile_pat_bindings_clocked_def: + compile_pat_bindings_clocked _ _ _ [] exp = (LN, exp) /\ + compile_pat_bindings_clocked 0 _ _ m exp = (LN, exp) /\ + compile_pat_bindings_clocked (SUC ck) t i ((Pany, _, _) :: m) exp = + compile_pat_bindings_clocked ck t i m exp /\ + compile_pat_bindings_clocked (SUC ck) t i ((Pvar s, k, x) :: m) exp = ( + let (spt, exp2) = compile_pat_bindings_clocked ck t i m exp in + (insert k () spt, Let t (SOME s) x exp2)) /\ + compile_pat_bindings_clocked (SUC ck) t i ((Plit _, _, _) :: m) exp = + compile_pat_bindings_clocked ck t i m exp /\ + compile_pat_bindings_clocked (SUC ck) t i ((Pcon _ ps, k, x) :: m) exp = ( + let j_nms = MAP (\(j, p). let k = i + 1 + j in + let nm = enc_num_to_name k [] in + ((j, nm), (p, k, Var_local t nm))) (enumerate 0 ps) in + let (spt, exp2) = compile_pat_bindings_clocked ck t (i + 2 + LENGTH ps) + (MAP SND j_nms ++ m) exp in + let j_nms_used = FILTER (\(_, (_, k, _)). IS_SOME (lookup k spt)) j_nms in + let exp3 = FOLDR (\((j, nm), _) exp. + flatLang$Let t (SOME nm) (App t (El j) [x]) exp) exp2 j_nms_used in + let spt2 = if NULL j_nms_used then spt else insert k () spt in + (spt2, exp3)) /\ + compile_pat_bindings_clocked (SUC ck) t i ((Pas p v, k, x) :: m) exp = ( + let nm = enc_num_to_name (i + 1) [] in + let (spt, exp2) = compile_pat_bindings_clocked ck t (i + 2) + ((p, i + 1, Var_local t nm) :: m) exp in + (insert k () spt, Let t (SOME v) x + (Let t (SOME nm) (Var_local t v) exp2))) /\ + compile_pat_bindings_clocked (SUC ck) t i ((Pref p, k, x) :: m) exp = ( + let nm = enc_num_to_name (i + 1) [] in + let (spt, exp2) = compile_pat_bindings_clocked ck t (i + 2) + ((p, i + 1, Var_local t nm) :: m) exp in + (insert k () spt, Let t (SOME nm) (App t (El 0) [x]) exp2)) +End + +val pre = cv_auto_trans_pre (compile_pat_bindings_clocked_def |> PURE_REWRITE_RULE[ELIM_UNCURRY]) + +Theorem compile_pat_bindings_clocked_pre[cv_pre,local]: + ∀v0 v1 v2 v exp. compile_pat_bindings_clocked_pre v0 v1 v2 v exp +Proof + ho_match_mp_tac compile_pat_bindings_clocked_ind >> + rw[] >> + rw[Once $ fetch "-" "compile_pat_bindings_clocked_pre_cases"] >> + gvs[ELIM_UNCURRY] +QED + +Theorem compile_pat_bindings_thm: + ∀n t i m exp. + SUM (MAP (pat_size o FST) m) + LENGTH m ≤ n ⇒ + compile_pat_bindings_clocked n t i m exp = compile_pat_bindings t i m exp +Proof + ho_match_mp_tac compile_pat_bindings_clocked_ind >> + rw[compile_pat_bindings_clocked_def,flat_patternTheory.compile_pat_bindings_def, + UNCURRY_eq_pair,PULL_EXISTS + ] + >- metis_tac[FST,SND,PAIR] + >> (PRED_ASSUM is_imp mp_tac >> + impl_tac + >- (gvs[flatLangTheory.pat_size_def,SUM_APPEND, flat_patternTheory.pat1_size, + LENGTH_enumerate, MAP_enumerate_MAPi, flat_patternTheory.MAPi_eq_MAP, + ADD1,o_DEF,MAP_MAP_o]) >> + strip_tac >> + gvs[] >> + metis_tac[FST,SND,PAIR]) +QED + +Theorem compile_pat_bindings_clocked_eq: + compile_pat_bindings t i m exp = compile_pat_bindings_clocked (SUM (MAP (pat_size o FST) m) + LENGTH m) t i m exp +Proof + irule $ GSYM compile_pat_bindings_thm >> rw[] +QED + +val _ = cv_auto_trans_pre flatLangTheory.pat_size_def + +Theorem flatLang_pat_size_pre[cv_pre]: + (∀v. flatLang_pat_size_pre v) ∧ + (∀v. flatLang_pat1_size_pre v) +Proof + Induct >> rw[Once $ fetch "-" "flatLang_pat_size_pre_cases"] +QED + +val _ = cv_auto_trans compile_pat_bindings_clocked_eq + +Definition naive_pattern_match_clocked_def: + naive_pattern_match_clocked 0 t xs = Bool t T /\ + naive_pattern_match_clocked (SUC ck) t [] = Bool t T /\ + naive_pattern_match_clocked (SUC ck) t ((flatLang$Pany, _) :: mats) = naive_pattern_match_clocked ck t mats + /\ + naive_pattern_match_clocked (SUC ck) t ((Pvar _, _) :: mats) = naive_pattern_match_clocked ck t mats /\ + naive_pattern_match_clocked (SUC ck) t ((Plit l, v) :: mats) = SmartIf t + (App t Equality [v; Lit t l]) (naive_pattern_match_clocked ck t mats) (Bool t F) /\ + naive_pattern_match_clocked (SUC ck) t ((Pcon NONE ps, v) :: mats) = + naive_pattern_match_clocked ck t (MAPi (\i p. (p, App t (El i) [v])) ps ++ mats) /\ + naive_pattern_match_clocked (SUC ck) t ((Pas p i, v) :: mats) = + naive_pattern_match_clocked ck t ((p, v) :: mats) /\ + naive_pattern_match_clocked (SUC ck) t ((Pcon (SOME stmp) ps, v) :: mats) = + SmartIf t (App t (TagLenEq (FST stmp) (LENGTH ps)) [v]) + (naive_pattern_match_clocked ck t (MAPi (\i p. (p, App t (El i) [v])) ps ++ mats)) + (Bool t F) + /\ + naive_pattern_match_clocked (SUC ck) t ((Pref p, v) :: mats) = + naive_pattern_match_clocked ck t ((p, App t (El 0) [v]) :: mats) +End + +val pre = cv_auto_trans_pre naive_pattern_match_clocked_def + +Theorem naive_pattern_match_clocked_pre[cv_pre]: + ∀v0 t v. naive_pattern_match_clocked_pre v0 t v +Proof + ho_match_mp_tac naive_pattern_match_clocked_ind >> + rw[] >> rw[Once $ fetch "-" "naive_pattern_match_clocked_pre_cases"] >> + gvs[MAPi_eq_list_mapi] +QED + +Theorem naive_pattern_match_clocked_thm: + ∀ck t xs. SUM (MAP (pat_size o FST) xs) + LENGTH xs ≤ ck ⇒ + naive_pattern_match_clocked ck t xs = naive_pattern_match t xs +Proof + ho_match_mp_tac naive_pattern_match_clocked_ind >> + rw[naive_pattern_match_clocked_def,flat_patternTheory.naive_pattern_match_def] >> + (PRED_ASSUM is_imp mp_tac >> + impl_tac + >- (gvs[flatLangTheory.pat_size_def,SUM_APPEND, flat_patternTheory.pat1_size, + LENGTH_enumerate, MAP_enumerate_MAPi, flat_patternTheory.MAPi_eq_MAP, + ADD1,o_DEF,MAP_MAP_o]) >> + strip_tac >> + gvs[] >> + metis_tac[FST,SND,PAIR]) +QED + +Theorem naive_pattern_matched_clocked_eq: + naive_pattern_match t xs = naive_pattern_match_clocked (SUM (MAP (pat_size o FST) xs) + LENGTH xs) t xs +Proof + irule $ GSYM naive_pattern_match_clocked_thm >> rw[] +QED + +val _ = cv_auto_trans naive_pattern_matched_clocked_eq + +Definition encode_pat_alt_def: + encode_pat_alt (flatLang$Pany) = pattern_semantics$Any /\ + encode_pat_alt (Plit l) = Lit l /\ + encode_pat_alt (Pvar _) = Any /\ + encode_pat_alt (flatLang$Pcon stmp ps) = Cons + (case stmp of NONE => NONE | SOME (i, NONE) => SOME (i, NONE) + | SOME (i, SOME (ty, ctors)) => SOME (i, SOME ctors)) + (encode_pats_alt ps) /\ + encode_pat_alt (Pas p v) = encode_pat_alt p /\ + encode_pat_alt (Pref p) = Ref (encode_pat_alt p) ∧ + encode_pats_alt [] = [] ∧ + encode_pats_alt (x::xs) = encode_pat_alt x::encode_pats_alt xs +End + +Theorem encode_pat_alt_thm: + (∀p. encode_pat_alt p = encode_pat p) ∧ + (∀ps. encode_pats_alt ps = MAP encode_pat ps) +Proof + Induct >> + rw[encode_pat_alt_def,flat_patternTheory.encode_pat_def] >> + metis_tac[] +QED + +val _ = cv_auto_trans encode_pat_alt_def + +val _ = cv_trans $ GSYM $ cj 1 encode_pat_alt_thm + +Definition exh_pat_alt_def: + exh_pat_alt Any = T /\ + exh_pat_alt (Or p1 p2) = (exh_pat_alt p1 \/ exh_pat_alt p2) /\ + exh_pat_alt (Cons NONE ps) = exh_pats_alt ps /\ + exh_pat_alt _ = F ∧ + exh_pats_alt [] = T ∧ + exh_pats_alt (x::xs) = (exh_pat_alt x ∧ exh_pats_alt xs) +Termination + wf_rel_tac ‘measure $ λx. sum_CASE x pat_size (list_size pat_size)’ +End + +val _ = cv_trans exh_pat_alt_def + +Theorem exh_pat_alt_thm: + (∀p. exh_pat_alt p = exh_pat p) ∧ + (∀ps. exh_pats_alt ps = EVERY exh_pat ps) +Proof + Induct >> + rw[exh_pat_alt_def,pattern_compTheory.exh_pat_def] >> + rename1 ‘Cons cc’ >> Cases_on ‘cc’ >> + rw[exh_pat_alt_def,pattern_compTheory.exh_pat_def] >> + metis_tac[] +QED + +val _ = cv_trans $ GSYM $ cj 1 exh_pat_alt_thm + +Definition sib_exists_alt_def: + sib_exists_alt [] t l = F ∧ + sib_exists_alt ((Cons (SOME (t1,_)) ps) :: xs) t l = + (if (t = t1 ∧ l = LENGTH ps) then T else sib_exists_alt xs t l) ∧ + sib_exists_alt _ _ _ = F +End + +val _ = cv_trans sib_exists_alt_def + +Theorem sib_exists_alt_thm: + ∀xs tl. sib_exists xs tl = sib_exists_alt xs (FST tl) (SND tl) +Proof + simp[FORALL_PROD] >> + recInduct sib_exists_alt_ind >> + rw[sib_exists_alt_def,pattern_compTheory.sib_exists_def] >> + metis_tac[] +QED + +val _ = cv_trans sib_exists_alt_thm + +val _ = cv_auto_trans pattern_compTheory.exh_rows_def + +val _ = cv_auto_trans pattern_compTheory.pat_to_guard_def + +val _ = cv_auto_trans flat_patternTheory.compile_pats_def + +val _ = cv_trans_rec flat_patternTheory.sum_string_ords_def + (wf_rel_tac ‘measure $ λ(x,y). cv_size(cv_LENGTH y) - cv_size x’ >> + cv_termination_tac >> + gvs[cvTheory.c2b_def,oneline cvTheory.cv_lt_def0,AllCaseEqs(), + oneline cvTheory.b2c_def]) + +val pre = cv_trans_pre flat_patternTheory.dec_name_to_num_def + +Theorem flat_pattern_dec_name_to_num_pre[cv_pre]: + ∀name. flat_pattern_dec_name_to_num_pre name +Proof + rw[fetch "-" "flat_pattern_dec_name_to_num_pre_cases"] >> + Cases_on ‘name’ >> gvs[] +QED + +Definition compile_exp_alt_def: + (compile_exp_alt cfg (Var_local t vid) = + (dec_name_to_num vid, F, Var_local t vid)) /\ + (compile_exp_alt cfg (Raise t x) = + let (i, sg, y) = compile_exp_alt cfg x in + (i, sg, Raise t y)) /\ + (compile_exp_alt cfg (Handle t x ps) = + let (i, sgx, y) = compile_exp_alt cfg x in + let (j, sgp, ps2) = compile_match_alt cfg ps in + let k = MAX i j + 2 in + let nm = enc_num_to_name k [] in + let v = Var_local t nm in + let r = Raise t v in + let exp = compile_pats cfg sgp t k v r ps2 in + (k, sgx \/ sgp, Handle t y [(Pvar nm, exp)])) /\ + (compile_exp_alt cfg (Con t ts xs) = + let (i, sg, ys) = compile_exp_alts_alt cfg (REVERSE xs) in + (i, sg, Con t ts (REVERSE ys))) /\ + (compile_exp_alt cfg (Fun t vs x) = + let (i, sg, y) = compile_exp_alt cfg x in + (i, sg, Fun t vs y)) /\ + (compile_exp_alt cfg (App t op xs) = + let (i, sg, ys) = compile_exp_alts_alt cfg (REVERSE xs) in + (i, sg \/ op_sets_globals op, App t op (REVERSE ys))) /\ + (compile_exp_alt cfg (Mat t x ps) = + let (i, sgx, y) = compile_exp_alt cfg x in + let (j, sgp, ps2) = compile_match_alt cfg ps in + let k = MAX i j + 2 in + let nm = enc_num_to_name k [] in + let v = Var_local t nm in + let r = Raise t (Con t (SOME (bind_tag, NONE)) []) in + let exp = compile_pats cfg sgp t k v r ps2 in + (k, sgx \/ sgp, Let t (SOME nm) y exp)) /\ + (compile_exp_alt cfg (Let t v x1 x2) = + let (i, sg1, y1) = compile_exp_alt cfg x1 in + let (j, sg2, y2) = compile_exp_alt cfg x2 in + let k = (case v of NONE => 0 | SOME vid => dec_name_to_num vid) in + (MAX i (MAX j k), sg1 \/ sg2, Let t v y1 y2)) /\ + (compile_exp_alt cfg (flatLang$Letrec t fs x) = + let ys = compile_letexps_alt cfg fs in + let (i, sgx, y) = compile_exp_alt cfg x in + let j = list_max (MAP (\(_,_,(j,_,_)). j) ys) in + let sgfs = EXISTS (\(_,_,(_,sg,_)). sg) ys in + let fs2 = MAP (\(a, b, (_, _, exp)). (a, b, exp)) ys in + (MAX i j, sgfs \/ sgx, flatLang$Letrec t fs2 y)) /\ + (compile_exp_alt cfg (If t x1 x2 x3) = + let (i, sg1, y1) = compile_exp_alt cfg x1 in + let (j, sg2, y2) = compile_exp_alt cfg x2 in + let (k, sg3, y3) = compile_exp_alt cfg x3 in + (MAX i (MAX j k), sg1 \/ sg2 \/ sg3, SmartIf t y1 y2 y3)) /\ + (compile_exp_alt cfg exp = (0, F, exp)) /\ + (compile_exp_alts_alt cfg [] = (0, F, [])) /\ + (compile_exp_alts_alt cfg (x::xs) = + let (i, sgx, y) = compile_exp_alt cfg x in + let (j, sgy, ys) = compile_exp_alts_alt cfg xs in + (MAX i j, sgx \/ sgy, y :: ys)) /\ + (compile_letexps_alt cfg [] = []) ∧ + (compile_letexps_alt cfg ((a,b,c)::xs) = + (a,b,compile_exp_alt cfg c)::compile_letexps_alt cfg xs) ∧ + (compile_match_alt cfg [] = (0, F, [])) /\ + (compile_match_alt cfg ((p, x)::ps) = + let (i, sgx, y) = compile_exp_alt cfg x in + let j = max_dec_name (pat_bindings p []) in + let (k, sgp, ps2) = compile_match_alt cfg ps in + (MAX i (MAX j k), sgx \/ sgp, ((p, y) :: ps2))) +Termination + WF_REL_TAC `measure (\x. case x of INL (_, x) => exp_size x + | INR (INL (_, xs)) => exp6_size xs + | INR (INR (INL (_, ps))) => exp1_size ps + | INR (INR (INR (_, ps))) => exp3_size ps)` + \\ rw [flatLangTheory.exp_size_def] +End + +(* TODO: move *) +Theorem cv_REV_size: + ∀x acc. cv_size(cv_REV x acc) ≤ cv_size x + cv_size acc +Proof + Induct_on ‘x’ >> + rw[] >> + rw[Once cv_stdTheory.cv_REV_def] >> + irule LESS_EQ_TRANS >> + first_x_assum $ irule_at $ Pos hd >> + rw[] +QED + +(* TODO: move *) +Theorem cv_REVERSE_size: + ∀x. cv_size(cv_REVERSE x) ≤ cv_size x +Proof + simp[cv_REVERSE_def] >> + Cases + >- (rw[Once cv_REV_def]) >> + irule_at (Pos hd) LESS_EQ_TRANS >> + irule_at (Pos hd) cv_REV_size >> + simp[] +QED + +val pre = cv_auto_trans_pre_rec compile_exp_alt_def + (WF_REL_TAC `measure (\x. case x of INL (_, x) => cv_size x + | INR (INL (_, xs)) => cv_size xs + | INR (INR (INL (_, ps))) => cv_size ps + | INR (INR (INR (_, ps))) => cv_size ps)` >> + cv_termination_tac >> + irule LESS_EQ_LESS_TRANS >> + irule_at (Pos last) cv_REVERSE_size >> + rw[]) + +Theorem compile_exp_alt_pre[cv_pre]: + (∀cfg v. compile_exp_alt_pre cfg v) ∧ + (∀cfg v. compile_exp_alts_alt_pre cfg v) ∧ + (∀cfg v. compile_letexps_alt_pre cfg v) ∧ + (∀cfg v. compile_match_alt_pre cfg v) +Proof + ho_match_mp_tac compile_exp_alt_ind >> + rw[] >> + rw[Once $ fetch "-" "compile_exp_alt_pre_cases"] +QED + +Theorem compile_exp_alt_thm: + (∀cfg exp. compile_exp_alt cfg exp = compile_exp cfg exp) ∧ + (∀cfg xs. compile_exp_alts_alt cfg xs = compile_exps cfg xs) /\ + (∀cfg xs. compile_letexps_alt cfg xs = MAP (λ(x,y,z). (x,y,compile_exp cfg z)) xs) ∧ + (∀cfg xs. compile_match_alt cfg xs = compile_match cfg xs) +Proof + ho_match_mp_tac compile_exp_alt_ind >> + rw[compile_exp_alt_def,flat_patternTheory.compile_exp_def] >> + rpt(pairarg_tac >> gvs[]) +QED + +val _ = cv_trans $ GSYM $ cj 1 compile_exp_alt_thm + +val _ = cv_trans flat_patternTheory.compile_dec_def + +(* flat_elim *) + +val _ = cv_auto_trans_pre flat_elimTheory.has_Eval_def; + +Theorem flat_elim_has_Eval_pre[cv_pre]: + (∀v. flat_elim_has_Eval_pre v) ∧ + (∀v. flat_elim_has_Eval_list_pre v) ∧ + (∀v. flat_elim_has_Eval_pats_pre v) ∧ + (∀v. flat_elim_has_Eval_funs_pre v) +Proof + ho_match_mp_tac flat_elimTheory.has_Eval_ind >> + rw[] >> rw[Once $ fetch "-" "flat_elim_has_Eval_pre_cases"] +QED + +Theorem cv_size_map_snd: + ∀z. cv_size(cv_map_snd z) ≤ cv_size z +Proof + Induct >> rw[] >> + rw[Once cv_stdTheory.cv_map_snd_def] >> + Cases_on ‘z’ >> rw[] +QED + +val pre = cv_auto_trans_pre_rec (flat_elimTheory.find_loc_def |> PURE_REWRITE_RULE[o_DEF,GSYM MAP_MAP_o]) + (WF_REL_TAC `measure (λ e . case e of + | INL x => cv_size x + | INR y => cv_size y)` >> + cv_termination_tac + >- (irule LESS_EQ_LESS_TRANS >> + irule_at (Pos last) cv_size_map_snd >> + rw[oneline cvTheory.cv_snd_def] >> + rpt(PURE_FULL_CASE_TAC >> gvs[])) + >- (irule LESS_EQ_LESS_TRANS >> + irule_at (Pos last) cv_size_map_snd >> + irule LESS_EQ_LESS_TRANS >> + irule_at (Pos last) cv_size_map_snd >> + rw[oneline cvTheory.cv_snd_def] >> + rpt(PURE_FULL_CASE_TAC >> gvs[])) + >- (irule LESS_EQ_LESS_TRANS >> + irule_at (Pos last) cv_size_map_snd >> + rw[oneline cvTheory.cv_snd_def] >> + rpt(PURE_FULL_CASE_TAC >> gvs[]))) + +Theorem flat_elim_find_loc_pre[cv_pre]: + (∀v. flat_elim_find_loc_pre v) ∧ + (∀v. flat_elim_find_locL_pre v) +Proof + ho_match_mp_tac flat_elimTheory.find_loc_ind >> + rw[] >> rw[Once $ fetch "-" "flat_elim_find_loc_pre_cases"] >> + rw[MAP_MAP_o] +QED + +val pre = cv_auto_trans_pre_rec (flat_elimTheory.find_lookups_def |> PURE_REWRITE_RULE[GSYM MAP_MAP_o,o_THM]) + (WF_REL_TAC `measure (λ e . case e of + | INL x => cv_size x + | INR y => cv_size y)` >> + cv_termination_tac + >- (irule LESS_EQ_LESS_TRANS >> + irule_at (Pos last) cv_size_map_snd >> + rw[oneline cvTheory.cv_snd_def] >> + rpt(PURE_FULL_CASE_TAC >> gvs[])) + >- (irule LESS_EQ_LESS_TRANS >> + irule_at (Pos last) cv_size_map_snd >> + irule LESS_EQ_LESS_TRANS >> + irule_at (Pos last) cv_size_map_snd >> + rw[oneline cvTheory.cv_snd_def] >> + rpt(PURE_FULL_CASE_TAC >> gvs[])) + >- (irule LESS_EQ_LESS_TRANS >> + irule_at (Pos last) cv_size_map_snd >> + rw[oneline cvTheory.cv_snd_def] >> + rpt(PURE_FULL_CASE_TAC >> gvs[]))) + +Theorem flat_elim_find_lookups_pre[cv_pre]: + (∀v. flat_elim_find_lookups_pre v) ∧ + (∀v. flat_elim_find_lookupsL_pre v) +Proof + ho_match_mp_tac flat_elimTheory.find_lookups_ind >> + rw[] >> rw[Once pre] >> gvs[GSYM MAP_MAP_o] +QED + +val _ = cv_auto_trans flat_elimTheory.total_pat_def; + +Definition is_pure_alt_def: + (is_pure_alt (Handle t e pes) = is_pure_alt e) ∧ + (is_pure_alt (Lit t l) = T) ∧ + (is_pure_alt (Con t id_option es) = is_pure_alts es) ∧ + (is_pure_alt (Var_local t str) = T) ∧ + (is_pure_alt (Fun t name body) = T) ∧ + (is_pure_alt (App t (GlobalVarInit g) es) = is_pure_alts es) ∧ + (is_pure_alt (If t e1 e2 e3) = (is_pure_alt e1 ∧ is_pure_alt e2 ∧ is_pure_alt e3)) ∧ + (is_pure_alt (Mat t e1 pes) = + (is_pure_alt e1 ∧ is_pure_alts (MAP SND pes) ∧ EXISTS total_pat (MAP FST pes))) ∧ + (is_pure_alt (Let t opt e1 e2) = (is_pure_alt e1 ∧ is_pure_alt e2)) ∧ + (is_pure_alt (Letrec t funs e) = is_pure_alt e) ∧ + (is_pure_alt _ = F) ∧ + (is_pure_alts [] = T) ∧ + (is_pure_alts (x::xs) = (is_pure_alt x ∧ is_pure_alts xs)) +Termination + WF_REL_TAC `measure (λ e . sum_CASE e exp_size $ list_size exp_size)` >> + rw[flatLangTheory.exp3_size] >> + ‘list_size exp_size (MAP SND pes) ≤ LENGTH pes + (SUM (MAP exp5_size pes))’ + suffices_by gvs[] >> + Induct_on ‘pes’ >> + rw[list_size_def,ADD1] >> + rename1 ‘SND xx’ >> Cases_on ‘xx’ >> rw[flatLangTheory.exp_size_def] +End + +val pre = cv_auto_trans_pre_rec is_pure_alt_def + (WF_REL_TAC `measure (λ e . sum_CASE e cv_size cv_size)` >> + cv_termination_tac >> + irule LESS_EQ_LESS_TRANS >> + irule_at (Pos last) cv_size_map_snd >> + rw[oneline cvTheory.cv_snd_def] >> + rpt(PURE_FULL_CASE_TAC >> gvs[])) + +Theorem is_pure_alt_pre[cv_pre]: + (∀v. is_pure_alt_pre v) ∧ + (∀v. is_pure_alts_pre v) +Proof + ho_match_mp_tac is_pure_alt_ind >> + rw[] >> rw[Once pre] +QED + +Theorem is_pure_alt_thm: + (∀v. is_pure_alt v = is_pure v) ∧ + (∀v. is_pure_alts v = EVERY is_pure v) +Proof + ho_match_mp_tac is_pure_alt_ind >> + rw[is_pure_alt_def,flat_elimTheory.is_pure_def] >> + metis_tac[] +QED + +val _ = cv_trans $ GSYM $ cj 1 is_pure_alt_thm + +Definition is_hidden_alt_def: + (is_hidden_alt (Raise t e) = is_hidden_alt e) ∧ + (is_hidden_alt (Handle t e pes) = F) ∧ + (is_hidden_alt (Lit t l) = T) ∧ + (is_hidden_alt (Con t id_option es) = is_hidden_alts es) ∧ + (is_hidden_alt (Var_local t str) = T) ∧ + (is_hidden_alt (Fun t name body) = T) ∧ + (is_hidden_alt (App t Opapp l) = F) ∧ + (is_hidden_alt (App t (GlobalVarInit g) [e]) = is_hidden_alt e) ∧ + (is_hidden_alt (App t (GlobalVarLookup g) [e]) = F) ∧ + (is_hidden_alt (If t e1 e2 e3) = (is_hidden_alt e1 ∧ is_hidden_alt e2 ∧ is_hidden_alt e3)) ∧ + (is_hidden_alt (Mat t e1 [p,e2]) = (is_hidden_alt e1 ∧ is_hidden_alt e2)) ∧ + (is_hidden_alt (Let t opt e1 e2) = (is_hidden_alt e1 ∧ is_hidden_alt e2)) ∧ + (is_hidden_alt (Letrec t funs e) = is_hidden_alt e) ∧ + (is_hidden_alt _ = F) ∧ + (is_hidden_alts [] = T) ∧ + (is_hidden_alts (x::xs) = (is_hidden_alt x ∧ is_hidden_alts xs)) +Termination + WF_REL_TAC `measure (λ e . sum_CASE e exp_size (list_size exp_size))` +End + +val pre = cv_trans_pre is_hidden_alt_def + +Theorem is_hidden_alt_pre[cv_pre]: + (∀v. is_hidden_alt_pre v) ∧ + (∀v. is_hidden_alts_pre v) +Proof + ho_match_mp_tac is_hidden_alt_ind >> + rw[] >> rw[Once $ fetch "-" "is_hidden_alt_pre_cases"] +QED + +Theorem is_hidden_alt_thm: + (∀v. is_hidden_alt v = is_hidden v) ∧ + (∀v. is_hidden_alts v = EVERY is_hidden v) +Proof + ho_match_mp_tac is_hidden_alt_ind >> + rw[is_hidden_alt_def,flat_elimTheory.is_hidden_def] >> + metis_tac[] +QED + +val _ = cv_trans $ GSYM $ cj 1 is_hidden_alt_thm + +Definition spt_fold_union_def: + (spt_fold_union acc LN = acc) ∧ + (spt_fold_union acc (LS a) = sptree$union a acc) ∧ + (spt_fold_union acc (BN t1 t2) = spt_fold_union (spt_fold_union acc t1) t2) ∧ + (spt_fold_union acc (BS t1 a t2) = spt_fold_union (union a $ spt_fold_union acc t1) t2) +End + +Theorem spt_fold_union_thm: + ∀acc t. spt_fold_union acc t = spt_fold union acc t +Proof + Induct_on ‘t’ >> rw[spt_fold_def,spt_fold_union_def] +QED + +val _ = cv_trans spt_fold_union_def + +val pre = cv_auto_trans_pre (spt_closureTheory.closure_spt_def |> PURE_REWRITE_RULE[GSYM spt_fold_union_thm]) + +Theorem spt_closure_closure_spt_pre[cv_pre]: + ∀reachable tree. spt_closure_closure_spt_pre reachable tree +Proof + ho_match_mp_tac spt_closureTheory.closure_spt_ind >> + rw[] >> rw[Once pre] >> + gvs[spt_fold_union_thm] +QED + +val _ = cv_auto_trans flat_elimTheory.remove_flat_prog_def; + +val _ = cv_auto_trans backend_asmTheory.to_flat_def; + +(* flat_to_clos *) + +val _ = cv_auto_trans flat_to_closTheory.compile_op_def + +Theorem list_size_thm: + list_size f xs = LENGTH xs + SUM(MAP f xs) +Proof + Induct_on ‘xs’ >> gvs[list_size_def] +QED + +Definition flat_to_clos_compile_alt_def: + (flat_to_clos_compile_alts m [] = []) /\ + (flat_to_clos_compile_alts m (x::xs) = flat_to_clos_compile_alt m x :: flat_to_clos_compile_alts m xs) /\ + (flat_to_clos_compile_alt m (flatLang$Raise t e) = (closLang$Raise t (flat_to_clos_compile_alt m (e)))) /\ + (flat_to_clos_compile_alt m (Lit t l) = (compile_lit t l)) /\ + (flat_to_clos_compile_alt m (Var_local t v) = (Var t (findi (SOME v) m))) /\ + (flat_to_clos_compile_alt m (Con t n es) = + let tag = (case n of SOME (t,_) => t | _ => 0) in + (SmartCons t tag (flat_to_clos_compile_alts m (REVERSE es)))) /\ + (flat_to_clos_compile_alt m (App t op es) = + case dest_nop op es of + | SOME e => flat_to_clos_compile_alt m e + | NONE => (compile_op t op (flat_to_clos_compile_alts m (REVERSE es)))) /\ + (flat_to_clos_compile_alt m (Fun t v e) = + (Fn (mlstring$implode t) NONE NONE 1 (flat_to_clos_compile_alt (SOME v::m) (e)))) /\ + (flat_to_clos_compile_alt m (If t x1 x2 x3) = + (If t (flat_to_clos_compile_alt m (x1)) + (flat_to_clos_compile_alt m (x2)) + (flat_to_clos_compile_alt m (x3)))) /\ + (flat_to_clos_compile_alt m (Let t vo e1 e2) = + (Let t [flat_to_clos_compile_alt m (e1)] (flat_to_clos_compile_alt (vo::m) (e2)))) /\ + (flat_to_clos_compile_alt m (Mat t e pes) = (Op t (Const 0) [])) /\ + (flat_to_clos_compile_alt m (Handle t e pes) = + case dest_pat pes of + | SOME (v,h) => (Handle t (flat_to_clos_compile_alt m (e)) (flat_to_clos_compile_alt (SOME v::m) (h))) + | _ => flat_to_clos_compile_alt m (e)) /\ + (flat_to_clos_compile_alt m (Letrec t funs e) = + let new_m = MAP (\n. SOME (FST n)) funs ++ m in + (Letrec (MAP (\n. join_strings (mlstring$implode t) (mlstring$implode (FST n))) funs) NONE NONE + (flat_to_clos_compile_lets_alt new_m funs) + (flat_to_clos_compile_alt new_m (e)))) ∧ + (flat_to_clos_compile_lets_alt m [] = []) /\ + (flat_to_clos_compile_lets_alt m ((f,v,x)::xs) = (1, flat_to_clos_compile_alt (SOME v :: m) x) :: flat_to_clos_compile_lets_alt m xs) +Termination + wf_rel_tac ‘measure $ λx. + case x of + INL (m, e) => list_size exp_size e + | INR (INL (m,e)) => exp_size e + | INR (INR (m,e)) => list_size (exp_size o SND o SND) e’ >> + rw[flatLangTheory.exp1_size,flatLangTheory.exp6_size, + list_size_thm + ] >> + gvs[oneline flat_to_closTheory.dest_pat_def,AllCaseEqs(), + oneline flat_to_closTheory.dest_nop_def, + flatLangTheory.op_size_def, + MAP_REVERSE,SUM_REVERSE,SUM_APPEND, + flatLangTheory.exp_size_def + ] >> + qmatch_goalsub_abbrev_tac ‘SUM (MAP a1 funs)’ >> + ‘SUM(MAP a1 funs) ≤ SUM (MAP exp2_size funs)’ + suffices_by gvs[] >> + unabbrev_all_tac >> + irule SUM_MAP_same_LE >> + simp[EVERY_MEM,FORALL_PROD] >> + rw[flatLangTheory.exp_size_def] +End + +val _ = cv_auto_trans flat_to_closTheory.dest_nop_def +val _ = cv_auto_trans flat_to_closTheory.dest_pat_def + +val pre = cv_auto_trans_pre_rec flat_to_clos_compile_alt_def + (wf_rel_tac ‘measure $ λx. + case x of + INL (m, e) => cv_size e + | INR (INL (m,e)) => cv_size e + | INR (INR (m,e)) => cv_size e’ >> + cv_termination_tac + >~ [‘cv_REVERSE’] + >- (irule LESS_EQ_LESS_TRANS >> + irule_at (Pos last) cv_REVERSE_size >> + rw[]) + >~ [‘cv_REVERSE’] + >- (irule LESS_EQ_LESS_TRANS >> + irule_at (Pos last) cv_REVERSE_size >> + rw[]) >> + gvs[fetch "-" "cv_flat_to_clos_dest_nop_def", + fetch "-" "cv_flat_to_clos_dest_pat_def", + AllCaseEqs()] >> + simp[oneline cvTheory.cv_snd_def, oneline cvTheory.cv_fst_def] >> + rpt(PURE_FULL_CASE_TAC >> gvs[])) + +Theorem flat_to_clos_compile_alts_pre[cv_pre]: + (∀m v. flat_to_clos_compile_alts_pre m v) ∧ + (∀m v. flat_to_clos_compile_alt_pre m v) ∧ + (∀m v. flat_to_clos_compile_lets_alt_pre m v) +Proof + ho_match_mp_tac flat_to_clos_compile_alt_ind >> + rw[] >> rw[Once pre] +QED + +Theorem flat_to_clos_compile_alt_thm: + (∀m xs. flat_to_clos_compile_alts m xs = compile m xs) ∧ + (∀m x. flat_to_clos_compile_alt m x = HD(compile m [x])) ∧ + (∀m xs. flat_to_clos_compile_lets_alt m xs = + (MAP ( \ (f,v,x). (1, HD (compile (SOME v :: m) [x]))) xs)) +Proof + ho_match_mp_tac flat_to_clos_compile_alt_ind >> + rw[flat_to_clos_compile_alt_def,flat_to_closTheory.compile_def] + >- (Cases_on ‘xs’ >> + gvs[flat_to_closTheory.compile_def,flat_to_closTheory.LENGTH_compile]) >> + rpt(PURE_TOP_CASE_TAC >> gvs[]) +QED + +val _ = cv_trans $ GSYM $ cj 1 flat_to_clos_compile_alt_thm + +val pre = cv_auto_trans_pre_rec closLangTheory.has_install_def + (wf_rel_tac ‘measure $ λx. sum_CASE x cv_size cv_size’ >> + cv_termination_tac >> + irule LESS_EQ_LESS_TRANS >> + irule_at (Pos last) cv_size_map_snd >> + rw[oneline cvTheory.cv_snd_def] >> + rpt(PURE_FULL_CASE_TAC >> gvs[])) + +Theorem closLang_has_install_pre[cv_pre]: + (∀v. closLang_has_install_pre v) ∧ + (∀v. closLang_has_install_list_pre v) +Proof + ho_match_mp_tac closLangTheory.has_install_ind >> + rw[] >> rw[Once pre] +QED + +val _ = cv_auto_trans flat_to_closTheory.compile_prog_def + +(* to_clos *) + +val _ = cv_trans backend_asmTheory.to_clos_def + +(* clos_mti *) + +val _ = cv_trans clos_mtiTheory.collect_args_def + +val _ = cv_trans clos_mtiTheory.collect_apps_def + +Definition intro_multi_alt_def: + (intro_multi_alt 0 max_app exp = (Var (SourceLoc 0 0 0 0) 0)) ∧ + (intro_multi_alt (SUC ck) max_app (closLang$Var t n) = (Var t n)) ∧ + (intro_multi_alt (SUC ck) max_app (If t e1 e2 e3) = + (If t ((intro_multi_alt ck max_app (e1))) + ((intro_multi_alt ck max_app (e2))) + ((intro_multi_alt ck max_app (e3))))) ∧ + (intro_multi_alt (SUC ck) max_app (Let t es e) = + (Let t (intro_multis_alt ck max_app es) ((intro_multi_alt ck max_app (e))))) ∧ + (intro_multi_alt (SUC ck) max_app (Raise t e) = + (Raise t ((intro_multi_alt ck max_app (e))))) ∧ + (intro_multi_alt (SUC ck) max_app (Handle t e1 e2) = + (Handle t ((intro_multi_alt ck max_app (e1))) ((intro_multi_alt ck max_app (e2))))) ∧ + (intro_multi_alt (SUC ck) max_app (Tick t e) = + (Tick t ((intro_multi_alt ck max_app (e))))) ∧ + (intro_multi_alt (SUC ck) max_app (Call t ticks n es) = + (Call t ticks n (intro_multis_alt ck max_app es))) ∧ + (intro_multi_alt (SUC ck) max_app (App t NONE e es) = + let (es', e') = collect_apps max_app es e in + (App t NONE ((intro_multi_alt ck max_app (e'))) (intro_multis_alt ck max_app es'))) ∧ + (intro_multi_alt (SUC ck) max_app (App t (SOME l) e es) = + (App t (SOME l) ((intro_multi_alt ck max_app (e))) (intro_multis_alt ck max_app es))) ∧ + (intro_multi_alt (SUC ck) max_app (Fn t NONE NONE num_args e) = + let (num_args', e') = collect_args max_app num_args e in + (Fn t NONE NONE num_args' ((intro_multi_alt ck max_app (e'))))) ∧ + (intro_multi_alt (SUC ck) max_app (Fn t loc fvs num_args e) = + (Fn t loc fvs num_args ((intro_multi_alt ck max_app (e))))) ∧ + (intro_multi_alt (SUC ck) max_app (Letrec t NONE NONE funs e) = + (Letrec t NONE NONE (intro_multi_collect_alt ck max_app funs) + ((intro_multi_alt ck max_app (e))))) ∧ + (intro_multi_alt (SUC ck) max_app (Letrec t (SOME loc) fvs funs e) = + (Letrec t (SOME loc) fvs funs ((intro_multi_alt ck max_app (e))))) ∧ + (intro_multi_alt (SUC ck) max_app (Letrec t NONE (SOME fvs) funs e) = + (Letrec t NONE (SOME fvs) funs ((intro_multi_alt ck max_app (e))))) ∧ + (intro_multi_alt (SUC ck) max_app (Op t op es) = + (Op t op (intro_multis_alt ck max_app es))) ∧ + (intro_multis_alt 0 max_app _ = []) ∧ + (intro_multis_alt (SUC ck) max_app [] = []) ∧ + (intro_multis_alt (SUC ck) max_app (e1::es) = + intro_multi_alt ck max_app e1 :: intro_multis_alt ck max_app es) ∧ + (intro_multi_collect_alt 0 max_app _ = []) ∧ + (intro_multi_collect_alt (SUC ck) max_app [] = []) ∧ + (intro_multi_collect_alt (SUC ck) max_app ((num_args,e)::fs) = + let (num_args',e') = collect_args max_app num_args e in + ((num_args', intro_multi_alt ck max_app e') :: intro_multi_collect_alt ck max_app fs)) +Termination + WF_REL_TAC `measure $ λx. case x of + INL(n,_,_) => n + | INR(INL(n,_,_)) => n + | INR(INR(n,_,_)) => n` +End + +Definition exp_size_alt_def: + exp_size_alt (closLang$Var x y) = 1:num ∧ + exp_size_alt (If a b c d) = 1 + exp_size_alt b + exp_size_alt c + exp_size_alt d ∧ + exp_size_alt (Let a es e) = 1 + exp_size_alt e + exp_sizes_alt es ∧ + exp_size_alt (Handle a0 a1 a2) = 1 + exp_size_alt a1 + exp_size_alt a2 ∧ + exp_size_alt (Raise a0 a1) = 1 + exp_size_alt a1 ∧ + exp_size_alt (Tick a0 a1) = 1 + exp_size_alt a1 ∧ + exp_size_alt (Call a0 a1 a2 a3) = 1 + exp_sizes_alt a3 ∧ + exp_size_alt (App a0 a1 a2 a3) = 1 + exp_size_alt a2 + exp_sizes_alt a3 ∧ + exp_size_alt (Fn a0 a1 a2 a3 a4) = 1 + exp_size_alt a4 ∧ + exp_size_alt (Letrec a0 a1 a2 a3 a4) = 1 + exp_sizes_alt (MAP SND a3) + exp_size_alt a4 ∧ + exp_size_alt (Op a0 a1 a2) = 1 + exp_sizes_alt a2 ∧ + exp_sizes_alt [] = 0 ∧ + exp_sizes_alt (x::xs) = 1 + exp_size_alt x + exp_sizes_alt xs +Termination + WF_REL_TAC `measure $ λx. case x of + INL e => exp_size e + | INR es => exp3_size es` >> + rpt strip_tac >> + sg ‘∀x. closLang$exp3_size(MAP SND x) ≤ exp1_size x’ + >- (Induct + >- rw[closLangTheory.exp_size_def] >> + Cases >> rw[closLangTheory.exp_size_def]) >> + irule LESS_EQ_LESS_TRANS >> + first_x_assum $ irule_at $ Pos last >> + rw[] +End + +val pre = cv_auto_trans_pre_rec exp_size_alt_def + (WF_REL_TAC `measure $ λx. case x of + INL e => cv_size e + | INR es => cv_size es` >> + cv_termination_tac >> + irule LESS_EQ_LESS_TRANS >> + irule_at (Pos last) cv_size_map_snd >> + rw[oneline cvTheory.cv_snd_def] >> + rpt(PURE_FULL_CASE_TAC >> gvs[])) + +Theorem exp_size_alt_pre[cv_pre]: + (∀v. exp_size_alt_pre v) ∧ + (∀v. exp_sizes_alt_pre v) +Proof + ho_match_mp_tac exp_size_alt_ind >> + rw[] >> rw[Once pre] +QED + +Theorem exp_sizes_alt_cons: + exp_sizes_alt (x::xs) = 1 + exp_size_alt x + exp_sizes_alt xs +Proof + Cases_on ‘xs’ >> rw[exp_size_alt_def] +QED + +Theorem exp_sizes_alt_append: + ∀xs ys. exp_sizes_alt (xs ++ ys) = exp_sizes_alt xs + exp_sizes_alt ys +Proof + Induct_on ‘xs’ >> rw[exp_sizes_alt_cons] >> rw[exp_size_alt_def] +QED + +val pre = cv_auto_trans_pre intro_multi_alt_def + +Theorem intro_multi_alt_pre[cv_pre]: + (∀v0 max_app v. intro_multi_alt_pre v0 max_app v) ∧ + (∀v0 max_app v. intro_multis_alt_pre v0 max_app v) ∧ + (∀v0 max_app v. intro_multi_collect_alt_pre v0 max_app v) +Proof + ho_match_mp_tac intro_multi_alt_ind >> + rw[] >> + rw[Once pre] +QED + +Theorem exp_size_alt_pos[simp]: + ∀e. exp_size_alt e = 0 ⇔ F +Proof + Cases >> rw[exp_size_alt_def] +QED + +Theorem exp_sizes_alt_pos[simp]: + ∀es. exp_sizes_alt es = 0 ⇔ es = [] +Proof + Cases >> rw[exp_size_alt_def] +QED + +Theorem collect_apps_size_alt: + ∀max_app args e args' e'. + collect_apps max_app args e = (args',e') ⇒ + exp_size_alt e' ≤ exp_size_alt e +Proof + recInduct clos_mtiTheory.collect_apps_ind >> + rw[exp_size_alt_def,clos_mtiTheory.collect_apps_def] >> + gvs[exp_size_alt_def] +QED + +Theorem collect_apps_sizes_alt: + ∀max_app args e args' e'. + collect_apps max_app args e = (args',e') ⇒ + exp_size_alt e' + exp_sizes_alt args' ≤ exp_size_alt e + exp_sizes_alt args +Proof + recInduct clos_mtiTheory.collect_apps_ind >> + rw[exp_size_alt_def,clos_mtiTheory.collect_apps_def] >> + gvs[exp_size_alt_def,exp_sizes_alt_append] +QED + +Theorem collect_args_size_alt: + ∀max_app args e args' e'. + collect_args max_app args e = (args',e') ⇒ + exp_size_alt e' ≤ exp_size_alt e +Proof + recInduct clos_mtiTheory.collect_args_ind >> + rw[exp_size_alt_def,clos_mtiTheory.collect_args_def] >> + gvs[exp_size_alt_def] +QED + +Theorem intro_multi_cons: + ∀e es. intro_multi max_app (e::es) = HD(intro_multi max_app [e])::intro_multi max_app es +Proof + Induct_on ‘es’ >> gvs[clos_mtiTheory.intro_multi_def,clos_mtiTheory.intro_multi_length] +QED + +Theorem intro_multi_thm: + (∀ck max_app e. exp_size_alt e ≤ ck ⇒ intro_multi_alt ck max_app e = HD(intro_multi max_app [e])) ∧ + (∀ck max_app es. exp_sizes_alt es ≤ ck ⇒ intro_multis_alt ck max_app es = intro_multi max_app es) ∧ + (∀ck max_app es. exp_sizes_alt(MAP SND es) ≤ ck ⇒ intro_multi_collect_alt ck max_app es = + MAP (\(num_args, e). + let (num_args', e') = collect_args max_app num_args e in + (num_args', HD (intro_multi max_app [e']))) + es) +Proof + ho_match_mp_tac intro_multi_alt_ind >> + rw[exp_size_alt_def,intro_multi_alt_def,clos_mtiTheory.intro_multi_def] >> + rpt(pairarg_tac >> gvs[]) >> + imp_res_tac collect_args_size_alt >> + imp_res_tac collect_apps_sizes_alt >> + gvs[] >> + CONV_TAC $ RHS_CONV $ PURE_ONCE_REWRITE_CONV [intro_multi_cons] >> + gvs[] +QED + +Theorem intro_multi_alt_eq: + intro_multi max_app es = intro_multis_alt (exp_sizes_alt es) max_app es +Proof + irule $ GSYM $ cj 2 intro_multi_thm >> + rw[] +QED + +val _ = cv_trans intro_multi_alt_eq + +val _ = cv_trans clos_mtiTheory.compile_def + +(* clos_number *) + +val pre = cv_trans_pre_rec clos_numberTheory.renumber_code_locs_def + (wf_rel_tac ‘measure $ λx. sum_CASE x (cv_size o SND) (cv_size o SND)’ >> + cv_termination_tac >> + irule LESS_EQ_LESS_TRANS >> + irule_at (Pos last) cv_size_map_snd >> + rw[oneline cvTheory.cv_snd_def] >> + rpt(PURE_FULL_CASE_TAC >> gvs[])) + +Theorem clos_number_renumber_code_locs_list_pre[cv_pre]: + (∀n v. clos_number_renumber_code_locs_list_pre n v) ∧ + (∀n v. clos_number_renumber_code_locs_pre n v) +Proof + ho_match_mp_tac clos_numberTheory.renumber_code_locs_ind >> + rw[] >> rw[Once pre] +QED + +(* clos_known *) + +Definition get_size_sc_aux_alt_def: + (get_size_sc_aux_alt n (Var t v) = n - 1) /\ + (get_size_sc_aux_alt n (If t x1 x2 x3) = + let n = n - 1 in if n = 0 then 0 else + let n = get_size_sc_aux_alt n x1 in if n = 0 then 0 else + let n = get_size_sc_aux_alt n x2 in if n = 0 then 0 else + get_size_sc_aux_alt n x3) /\ + (get_size_sc_aux_alt n (Let t xs x2) = + let n = n - 1 in if n = 0 then 0 else + let n = get_size_sc_aux_alts n xs in if n = 0 then 0 else + get_size_sc_aux_alt n x2) /\ + (get_size_sc_aux_alt n (Raise t x1) = + let n = n - 1 in if n = 0 then 0 else + get_size_sc_aux_alt n x1) /\ + (get_size_sc_aux_alt n (Handle t x1 x2) = + let n = n - 1 in if n = 0 then 0 else + let n = get_size_sc_aux_alt n x1 in if n = 0 then 0 else + get_size_sc_aux_alt n x2) /\ + (get_size_sc_aux_alt n (Op t op xs) = + let n = n - 1 in if n = 0 then 0 else + get_size_sc_aux_alts n xs) /\ + (get_size_sc_aux_alt n (Tick t x) = get_size_sc_aux_alt n x) /\ + (get_size_sc_aux_alt n (Call t ticks dest xs) = + let n = n - 1 in if n = 0 then 0 else + get_size_sc_aux_alts n xs) /\ + (get_size_sc_aux_alt n (Fn t loc_opt ws_opt num_args x1) = + let n = n - 1 in if n = 0 then 0 else + get_size_sc_aux_alt n x1) /\ + (get_size_sc_aux_alt n (Letrec t loc_opt ws_opt fns x1) = + let n = n - 1 in if n = 0 then 0 else + let n = get_size_sc_aux_alts n (MAP SND fns) in if n = 0 then 0 else + get_size_sc_aux_alt n x1) /\ + (get_size_sc_aux_alt n (App t loc_opt x1 xs) = + let n = n - 1 in if n = 0 then 0 else + let n = get_size_sc_aux_alt n x1 in if n = 0 then 0 else + get_size_sc_aux_alts n xs) ∧ + (get_size_sc_aux_alts n [] = n) /\ + (get_size_sc_aux_alts n (x::xs) = + if n = 0n then n else + let n = get_size_sc_aux_alt n x in if n = 0 then n else + get_size_sc_aux_alts n xs) +Termination + WF_REL_TAC `measure $ λx. sum_CASE x (exp_size o SND) (exp3_size o SND)` + \\ simp [] \\ rpt strip_tac + \\ `exp3_size (MAP SND fns) <= exp1_size fns` + by (Induct_on `fns` \\ simp [closLangTheory.exp_size_def] \\ Cases \\ simp [closLangTheory.exp_size_def]) + \\ simp [] +End + +val pre = cv_auto_trans_pre_rec get_size_sc_aux_alt_def + (WF_REL_TAC `measure $ λx. sum_CASE x (cv_size o SND) (cv_size o SND)` >> + cv_termination_tac >> + irule LESS_EQ_LESS_TRANS >> + irule_at (Pos last) cv_size_map_snd >> + rw[oneline cvTheory.cv_snd_def] >> + rpt(PURE_FULL_CASE_TAC >> gvs[])) + +Theorem get_size_sc_aux_alt_pre[cv_pre]: + (∀n v. get_size_sc_aux_alt_pre n v) ∧ + (∀n v. get_size_sc_aux_alts_pre n v) +Proof + ho_match_mp_tac get_size_sc_aux_alt_ind >> + rw[] >> rw[Once pre] +QED + +Theorem get_size_aux_sc_alt_thm: + (∀n v. get_size_sc_aux_alt n v = get_size_sc_aux n [v]) ∧ + (∀n vs. get_size_sc_aux_alts n vs = get_size_sc_aux n vs) +Proof + ho_match_mp_tac get_size_sc_aux_alt_ind >> + rw[get_size_sc_aux_alt_def,clos_knownTheory.get_size_sc_aux_def] >> + gvs[] + >- gvs[clos_knownTheory.get_size_sc_aux_correct] >> + Cases_on ‘vs’ >> gvs[clos_knownTheory.get_size_sc_aux_def] +QED + +val _ = cv_trans $ GSYM $ cj 2 get_size_aux_sc_alt_thm + +val _ = cv_trans $ GSYM $ cj 2 get_size_aux_sc_alt_thm + +val _ = cv_trans clos_knownTheory.get_size_sc_def + +Definition free_alt_def: + (free_alt (closLang$Var t v) = ((closLang$Var t v), db_vars$Var v)) /\ + (free_alt (If t x1 x2 x3) = + let (c1,l1) = free_alt (x1) in + let (c2,l2) = free_alt (x2) in + let (c3,l3) = free_alt (x3) in + ((If t ( c1) ( c2) ( c3)),mk_Union l1 (mk_Union l2 l3))) /\ + (free_alt (Let t xs x2) = + let (c1,l1) = free_alts xs in + let (c2,l2) = free_alt (x2) in + ((Let t c1 ( c2)),mk_Union l1 (Shift (LENGTH xs) l2))) /\ + (free_alt (Raise t x1) = + let (c1,l1) = free_alt (x1) in + ((Raise t ( c1)),l1)) /\ + (free_alt (Tick t x1) = + let (c1,l1) = free_alt (x1) in + ((Tick t ( c1)),l1)) /\ + (free_alt (Op t op xs) = + let (c1,l1) = free_alts xs in + ((Op t op c1),l1)) /\ + (free_alt (App t loc_opt x1 xs2) = + let (c1,l1) = free_alt (x1) in + let (c2,l2) = free_alts xs2 in + ((App t loc_opt ( c1) c2),mk_Union l1 l2)) /\ + (free_alt (Fn t loc _ num_args x1) = + let (c1,l1) = free_alt (x1) in + let l2 = Shift num_args l1 in + ((Fn t loc (SOME (vars_to_list l2)) num_args ( c1)),l2)) /\ + (free_alt (Letrec t loc _ fns x1) = + let m = LENGTH fns in + let res = free_let_alts m fns in + let c1 = MAP FST res in + let l1 = list_mk_Union (MAP SND res) in + let (c2,l2) = free_alt (x1) in + ((Letrec t loc (SOME (vars_to_list l1)) c1 ( c2)), + mk_Union l1 (Shift (LENGTH fns) l2))) /\ + (free_alt (Handle t x1 x2) = + let (c1,l1) = free_alt (x1) in + let (c2,l2) = free_alt (x2) in + ((Handle t ( c1) ( c2)),mk_Union l1 (Shift 1 l2))) /\ + (free_alt (Call t ticks dest xs) = + let (c1,l1) = free_alts xs in + ((Call t ticks dest c1),l1)) ∧ + (free_alts [] = ([],Empty)) /\ + (free_alts ((x:closLang$exp)::xs) = + let (c1,l1) = free_alt x in + let (c2,l2) = free_alts xs in + (c1::c2,mk_Union l1 l2)) ∧ + (free_let_alts m [] = []) /\ + (free_let_alts m ((n,x)::xs) = + let (c,l) = free_alt (x) + in + ((n, c),Shift (n + m) l)::free_let_alts m xs) +Termination + WF_REL_TAC `measure $ λa. case a of + | INL x => closLang$exp_size x + | INR(INL xs) => exp3_size xs + | INR(INR(_,xs)) => exp3_size (MAP SND xs)` >> + rw[] >> gvs[closLangTheory.exp_size_eq,list_size_def] >> + ‘list_size exp_size (MAP SND fns) ≤ list_size (pair_size (λx. x) exp_size) fns’ + suffices_by simp[] >> + Induct_on ‘fns’ >> + simp[list_size_def] >> + Cases >> + rw[] >> + rw[basicSizeTheory.pair_size_def] +End + +val pre = cv_auto_trans_pre free_alt_def + +Theorem free_alt_pre[cv_pre]: + (∀v. free_alt_pre v) ∧ + (∀v. free_alts_pre v) ∧ + (∀m v. free_let_alts_pre m v) +Proof + ho_match_mp_tac free_alt_ind >> + rw[] >> rw[Once pre] +QED + +Triviality free_LENGTH_LEMMA: + !xs ys s1. free xs = (ys,s1) ⇒ LENGTH xs = LENGTH ys +Proof + recInduct clos_knownTheory.free_ind \\ rpt strip_tac + \\ gvs [clos_knownTheory.free_def,UNCURRY_eq_pair] +QED + +Theorem free_alt_thm: + (∀v. free_alt v = (λ(x,y). (HD x, y)) $ free [v]) ∧ + (∀v. free_alts v = free v) ∧ + (∀m v. free_let_alts m v = + MAP (\(n,x). let (c,l) = free [x] in + ((n,HD c),Shift (n + m) l)) v) +Proof + ho_match_mp_tac free_alt_ind >> + rw[free_alt_def,clos_knownTheory.free_def] >> + rpt(pairarg_tac >> gvs[]) >> + imp_res_tac free_LENGTH_LEMMA >> + gvs[LENGTH_EQ_1] >> + rename1 ‘free (_::xx)’ >> Cases_on ‘xx’ >> gvs[] >> + gvs[clos_knownTheory.free_def] +QED + +val _ = cv_trans $ GSYM $ cj 2 free_alt_thm + +val _ = cv_auto_trans clos_knownTheory.closed_def + +Definition contains_closures_alt_def: + (contains_closures_alt (closLang$Var t v) = F) /\ + (contains_closures_alt (If t x1 x2 x3) = + if contains_closures_alt (x1) then T else + if contains_closures_alt (x2) then T else + contains_closures_alt (x3)) /\ + (contains_closures_alt (Let t xs x2) = + if contains_closures_alts xs then T else + contains_closures_alt (x2)) /\ + (contains_closures_alt (Raise t x1) = contains_closures_alt (x1)) /\ + (contains_closures_alt (Handle t x1 x2) = + if contains_closures_alt (x1) then T else + contains_closures_alt (x2)) /\ + (contains_closures_alt (Op t op xs) = contains_closures_alts xs) /\ + (contains_closures_alt (Tick t x) = contains_closures_alt (x)) /\ + (contains_closures_alt (Call t ticks dest xs) = contains_closures_alts xs) /\ + (contains_closures_alt (Fn t loc_opt ws_opt num_args x1) = T) /\ + (contains_closures_alt (Letrec t loc_opt ws_opt fns x1) = T) /\ + (contains_closures_alt (App t loc_opt x1 xs) = + if contains_closures_alt (x1) then T else + contains_closures_alts xs) ∧ + (contains_closures_alts [] = F) /\ + (contains_closures_alts (x::xs) = + if contains_closures_alt x then T else + contains_closures_alts xs) +End + +val pre = cv_auto_trans_pre contains_closures_alt_def + +Theorem contains_closures_alt_pre[cv_pre]: + (∀v. contains_closures_alt_pre v) ∧ + (∀v. contains_closures_alts_pre v) +Proof + Induct >> rw[] >> rw[Once pre] +QED + +Theorem contains_closures_cons: + ∀v vs. contains_closures(v::vs) = (contains_closures [v] ∨ contains_closures vs) +Proof + Induct_on ‘vs’ >> + gvs[clos_knownTheory.contains_closures_def] +QED + +Theorem contains_closures_alt_thm: + (∀v. contains_closures_alt v = contains_closures [v]) ∧ + (∀v. contains_closures_alts v = contains_closures v) +Proof + Induct >> rw[contains_closures_alt_def,clos_knownTheory.contains_closures_def] >> + metis_tac[contains_closures_cons] +QED + +val _ = cv_trans $ GSYM $ cj 2 contains_closures_alt_thm + +Definition merge_alt_def: + (merge_alt Impossible y = y) ∧ + (merge_alt x Impossible = x) ∧ + (merge_alt (Tuple tg1 xs) (Tuple tg2 ys) = + if LENGTH xs = LENGTH ys ∧ tg1 = tg2 then Tuple tg1 (merge_alts xs ys) + else Other) ∧ + (merge_alt (ClosNoInline m1 n1) (ClosNoInline m2 n2) = + if m1 = m2 ∧ n1 = n2 + then ClosNoInline m1 n1 else Other) ∧ + (merge_alt (Clos m1 n1 e1 s1) (Clos m2 n2 e2 s2) = + if m1 = m2 ∧ n1 = n2 /\ e1 = e2 /\ s1 = s2 + then Clos m1 n1 e1 s1 else Other) ∧ + (merge_alt (Int i) (Int j) = if i = j then Int i else Other) ∧ + (merge_alt _ _ = Other) ∧ + (merge_alts (x::xs) (y::ys) = + merge_alt x y::merge_alts xs ys) ∧ + (merge_alts _ _ = []) +Termination + wf_rel_tac ‘measure $ λx. sum_CASE x (val_approx_size o FST) (val_approx1_size o FST)’ +End + +val pre = cv_auto_trans_pre merge_alt_def + +Theorem merge_alt_pre[cv_pre]: + (∀v0 v. merge_alt_pre v0 v) ∧ + (∀v0 v. merge_alts_pre v0 v) +Proof + ho_match_mp_tac merge_alt_ind >> + rw[] >> rw[Once pre] +QED + +Theorem merge_alt_thm: + (∀v0 v. merge_alt v0 v = merge v0 v) ∧ + (∀v0 v. merge_alts v0 v = MAP2 merge v0 v) +Proof + ho_match_mp_tac merge_alt_ind >> + rw[merge_alt_def,clos_knownTheory.merge_def] +QED + +val _ = cv_trans $ GSYM $ cj 1 merge_alt_thm + +val pre = cv_trans_pre clos_knownTheory.known_op_def + +Theorem clos_known_known_op_pre[cv_pre]: + ∀v as g. clos_known_known_op_pre v as g +Proof + ho_match_mp_tac clos_knownTheory.known_op_ind >> + rw[] >> rw[Once pre] >> + intLib.COOPER_TAC +QED + +val _ = cv_trans clos_knownTheory.clos_approx_def + +val _ = cv_trans clos_knownTheory.clos_gen_noinline_def + +val _ = cv_trans clos_knownTheory.isGlobal_def + +val _ = cv_trans clos_knownTheory.gO_destApx_def + +val _ = cv_trans clos_knownTheory.mk_Ticks_def + +Definition pure_alt_def: + (pure_alt (Var _ _) ⇔ T) + ∧ + (pure_alt (If _ e1 e2 e3) ⇔ pure_alt e1 ∧ pure_alt e2 ∧ pure_alt e3) + ∧ + (pure_alt (Let _ es e2) ⇔ pure_alts es ∧ pure_alt e2) + ∧ + (pure_alt (Raise _ _) ⇔ F) + ∧ + (pure_alt (Handle _ e1 e2) ⇔ pure_alt e1) + ∧ + (pure_alt (Tick _ _) ⇔ F) + ∧ + (pure_alt (Call _ _ _ _) ⇔ F) + ∧ + (pure_alt (App _ _ _ _) ⇔ F) + ∧ + (pure_alt (Fn _ _ _ _ _) ⇔ T) + ∧ + (pure_alt (Letrec _ _ _ _ x) ⇔ pure_alt x) + ∧ + (pure_alt (Op _ opn es) ⇔ pure_alts es ∧ pure_op opn) + ∧ + (pure_alts [] = T) + ∧ + (pure_alts (x::xs) ⇔ pure_alt x ∧ pure_alts xs) +End + +val pre = cv_auto_trans_pre pure_alt_def + +Theorem pure_alt_pre[cv_pre]: + (∀v. pure_alt_pre v) ∧ + (∀v. pure_alts_pre v) +Proof + Induct >> + rw[] >> rw[Once pre] +QED + +Theorem pure_alt_thm: + (∀v. pure_alt v = pure v) ∧ + (∀v. pure_alts v = EVERY pure v) +Proof + Induct >> rw[pure_alt_def,closLangTheory.pure_def] >> + metis_tac[] +QED + +val _ = cv_trans $ GSYM $ cj 1 pure_alt_thm + +Definition decide_inline_alt_def: + decide_inline_alt c fapx app_lopt app_arity = + case fapx of + | ClosNoInline loc arity => + if app_lopt = NONE /\ app_arity = arity + then inlD_Annotate loc + else inlD_Nothing + | Clos loc arity body body_size => + if app_lopt = NONE /\ app_arity = arity then + (if body_size < c * (1 + app_arity) /\ + ~contains_closures [body] /\ + closed (Fn (strlit "") NONE NONE app_arity body) + (* Consider moving these checks to the point where Clos approximations + are created, and bake them into the val_approx_val relation. *) + then inlD_LetInline body + else inlD_Annotate loc) + else inlD_Nothing + | _ => inlD_Nothing +End + +val _ = cv_auto_trans decide_inline_alt_def + +Theorem decide_inline_alt_thm: + decide_inline c fapx app_lopt app_arity = decide_inline_alt c.inline_factor fapx app_lopt app_arity +Proof + rw[clos_knownTheory.decide_inline_def,decide_inline_alt_def] +QED + +Theorem decide_inline_alt_LetInline: + !c fapx lopt arity body. + decide_inline_alt c fapx lopt arity = inlD_LetInline body ==> 0 < c +Proof + rpt strip_tac + \\ Cases_on `fapx` \\ fs [decide_inline_alt_def, bool_case_eq] + \\ spose_not_then assume_tac \\ fs [] +QED + +Definition known_alt_def: + (known_alt inl_factor inl_max (Var t v) vs g = + (((Var t v,any_el v vs Other)),g)) /\ + (known_alt inl_factor inl_max (If t x1 x2 x3) vs g = + let (ea1,g) = known_alt inl_factor inl_max (x1) vs g in + let (ea2,g) = known_alt inl_factor inl_max (x2) vs g in + let (ea3,g) = known_alt inl_factor inl_max (x3) vs g in + let (e1,a1) = ea1 in + let (e2,a2) = ea2 in + let (e3,a3) = ea3 in + (((If t e1 e2 e3), merge a2 a3),g)) /\ + (known_alt inl_factor inl_max (Let t xs x2) vs g = + let (ea1,g) = known_alts inl_factor inl_max xs vs g in + let (ea2,g) = known_alt inl_factor inl_max (x2) (MAP SND ea1 ++ vs) g in + let (e2,a2) = ea2 in + (((Let t (MAP FST ea1) e2, a2)),g)) /\ + (known_alt inl_factor inl_max (Raise t x1) vs g = + let (ea1,g) = known_alt inl_factor inl_max (x1) vs g in + let (e1,a1) = ea1 in + (((Raise t e1,Impossible)),g)) /\ + (known_alt inl_factor inl_max (Tick t x1) vs g = + let (ea1,g) = known_alt inl_factor inl_max (x1) vs g in + let (e1,a1) = ea1 in + (((Tick t e1,a1)),g)) /\ + (known_alt inl_factor inl_max (Handle t x1 x2) vs g = + let (ea1,g) = known_alt inl_factor inl_max (x1) vs g in + let (ea2,g) = known_alt inl_factor inl_max (x2) (Other::vs) g in + let (e1,a1) = ea1 in + let (e2,a2) = ea2 in + (((Handle t e1 e2,merge a1 a2)),g)) /\ + (known_alt inl_factor inl_max (Call t ticks dest xs) vs g = + let (ea1,g) = known_alts inl_factor inl_max xs vs g in + (((Call t ticks dest (MAP FST ea1),Other)),g)) /\ + (known_alt inl_factor inl_max (Op t op xs) vs g = + let (ea1,g) = known_alts inl_factor inl_max xs vs g in + let (a,g) = known_op op (REVERSE (MAP SND ea1)) g in + let e = + (if isGlobal op then + case gO_destApx a of + | gO_None => SmartOp t op (MAP FST ea1) + | gO_Int i => Op t (Const i) (MAP FST ea1) + | gO_NullTuple tag => Op t (Cons tag) (MAP FST ea1) + else SmartOp t op (MAP FST ea1)) + in + (((e,a)),g)) /\ + (known_alt inl_factor inl_max (App t loc_opt x xs) vs g = + let (ea2,g) = known_alts inl_factor inl_max xs vs g in + let (ea1,g) = known_alt inl_factor inl_max (x) vs g in + let (e1,a1) = ea1 + in + case decide_inline_alt inl_factor a1 loc_opt (LENGTH xs) of + | inlD_Nothing => (((App t loc_opt e1 (MAP FST ea2), Other)), g) + | inlD_Annotate new_loc => (((App t (SOME new_loc) e1 (MAP FST ea2), Other)), g) + | inlD_LetInline body => + let (eabody,_) = known_alt (inl_factor DIV 2) inl_max (body) (MAP SND ea2) g in + let (ebody, abody) = eabody in + if pure x then + (* We don't have to evaluate x *) + (((Let (t§0) (MAP FST ea2) + (mk_Ticks t 1 (LENGTH xs) ebody), abody)), g) + else + (* We need to evaluate x for its side-effects, + but we must do so after evaluating the args. *) + (((Let (t§0) (SNOC e1 (MAP FST ea2)) + (mk_Ticks t 1 (LENGTH xs) ebody), abody)), g)) /\ + (known_alt inl_factor inl_max (Fn t loc_opt ws_opt num_args x1) vs g = + let (ea1,g) = known_alt inl_factor inl_max (x1) (REPLICATE num_args Other ++ vs) g in + let (body,a1) = ea1 in + (((Fn t loc_opt NONE num_args body, + case loc_opt of + | SOME loc => clos_approx inl_max loc num_args x1 + | NONE => Other)), g)) /\ + (known_alt inl_factor inl_max (Letrec t loc_opt _ fns x1) vs g = + let clos = case loc_opt of + NONE => REPLICATE (LENGTH fns) Other + | SOME loc => clos_gen_noinline loc 0 fns in + (* The following ignores SetGlobal within fns, but it shouldn't + appear there, and missing it just means this opt will do less. *) + let new_fns = known_let_alts clos vs inl_factor inl_max g fns in + let (ea1,g) = known_alt inl_factor inl_max (x1) (clos ++ vs) g in + let (e1,a1) = ea1 in + (((Letrec t loc_opt NONE new_fns e1,a1)),g)) ∧ + (known_alts inl_factor inl_max [] vs (g:val_approx spt) = ([],g)) /\ + (known_alts inl_factor inl_max ((x:closLang$exp)::xs) vs g = + let (eas1,g) = known_alt inl_factor inl_max (x) vs g in + let (eas2,g) = known_alts inl_factor inl_max (xs) vs g in + (eas1::eas2,g)) ∧ + (known_let_alts _ _ _ _ _ [] = []) ∧ + (known_let_alts clos vs inl_factor inl_max g ((num_args,x)::xs) = + let + new_vs = REPLICATE num_args Other ++ clos ++ vs; + res = known_alt inl_factor inl_max x new_vs g + in + (num_args,FST (FST res))::known_let_alts clos vs inl_factor inl_max g xs) +Termination + wf_rel_tac `inv_image (measure I LEX measure I) + (λx. case x of + INL (c,_,x,vs,g) => (c, closLang$exp_size x) + | INR(INL (c,_,xs,vs,g)) => (c, closLang$exp3_size xs) + | INR(INR (clos,vs,c,_,g,xs)) => (c, closLang$exp1_size xs))` + \\ simp [clos_knownTheory.dec_inline_factor_def] \\ rpt strip_tac + \\ imp_res_tac decide_inline_alt_LetInline \\ fs [] +End + +val pre = cv_auto_trans_pre_rec clos_opTheory.lift_exps_def + (fn x => + (wf_rel_tac ‘measure $ λx. cv_size(FST x)’ >> + cv_termination_tac >> + gvs[fetch "-" "cv_dest_Op_dest_Cons_def",AllCaseEqs(), + cvTheory.c2b_def,cvTheory.cv_lt_def0, + cvTheory.b2c_if, + oneline cvTheory.cv_ispair_def, + fetch "-" "cv_clos_op_dest_Cons_def" + ] >> + gvs[oneline cvTheory.cv_fst_def, + oneline cvTheory.cv_snd_def, + AllCaseEqs()]) + x) + +Theorem clos_op_lift_exps_pre[cv_pre]: + ∀v n binds. clos_op_lift_exps_pre v n binds +Proof + ho_match_mp_tac clos_opTheory.lift_exps_ind >> + rw[] >> rw[Once pre] +QED + +Definition eq_pure_list_alt: + eq_pure_list_alt (SUC ck) [(x,y)] = + (case eq_direct x y of + | SOME z => List [z] + | NONE => + case dest_Op dest_Cons x, dest_Op dest_Cons y of + | (NONE, NONE) => List [Op None Equal [x;y]] + | (SOME (t1,xs), SOME (t2,ys)) => + if t1 ≠ t2 ∨ LENGTH xs ≠ LENGTH ys then List [MakeBool F] + else eq_pure_list_alt ck (ZIP (REVERSE xs, REVERSE ys)) + | (SOME (t1,xs), NONE) => + Append (List [Op None (TagLenEq t1 (LENGTH xs)) [y]]) + (eq_pure_list_alt ck (MAPi (λi x. (x, Op None (ElemAt i) [y])) (REVERSE xs))) + | (NONE, SOME (t1,ys)) => eq_pure_list_alt ck [(y,x)]) ∧ + eq_pure_list_alt (SUC ck) (xy::xys) = Append (eq_pure_list_alt ck [xy]) (eq_pure_list_alt ck xys) ∧ + eq_pure_list_alt _ _ = Nil +End + +val _ = cv_auto_trans_pre eq_pure_list_alt + +Theorem eq_pure_list_alt_pre[cv_pre]: + ∀v0 v. eq_pure_list_alt_pre v0 v +Proof + ho_match_mp_tac eq_pure_list_alt_ind >> rw[] >> + rw[Once $ fetch "-" "eq_pure_list_alt_pre_cases"] >> + rw[] >> + gvs[cv_stdTheory.MAPi_eq_list_mapi] +QED + +Triviality SUM_MAP_const = + SUM_MAP_PLUS + |> CONV_RULE SWAP_FORALL_CONV + |> Q.SPEC ‘K x’ + |> SIMP_RULE std_ss [MAP_K_REPLICATE,SUM_REPLICATE]; + +Theorem eq_pure_list_alt_thm: + ∀ck xs. SUM (MAP (λ(x,y). 1 + cons_measure x + cons_measure y + + if cons_measure x < cons_measure y then 1 else 0) xs) ≤ ck ⇒ eq_pure_list_alt ck xs = eq_pure_list xs +Proof + recInduct eq_pure_list_alt_ind >> rpt strip_tac >> + gvs[] >> + rw[eq_pure_list_alt,Once clos_opTheory.eq_pure_list_def] >> + rpt(IF_CASES_TAC ORELSE PURE_FULL_CASE_TAC >> gvs[GSYM cv_stdTheory.MAPi_eq_list_mapi]) >> + imp_res_tac clos_opTheory.cons_measure_lemma >> + gvs[UNCURRY_eq_pair] >> + rpt(pairarg_tac >> gvs[]) + >- (first_x_assum irule >> + gvs[o_DEF] >> + gvs[clos_opTheory.cons_measure_lemma,MAP_REVERSE,SUM_REVERSE,SUM_MAP_const]) + >- (first_x_assum irule >> + gvs[GSYM REVERSE_ZIP,MAP_REVERSE,SUM_REVERSE] >> + rename1 ‘ZIP (xs, ys)’ >> + qmatch_goalsub_abbrev_tac ‘MAP f _’ >> + irule LESS_EQ_TRANS >> + Q.SUBGOAL_THEN + ‘SUM (MAP f (ZIP (xs, ys))) + ≤ SUM(MAP (λx. cons_measure x + 1) xs) + SUM(MAP (λx. cons_measure x + 1) ys)’ + $ irule_at $ Pos hd >> + unabbrev_all_tac + >- (qpat_x_assum ‘LENGTH _ = LENGTH _’ mp_tac >> + rpt $ pop_assum kall_tac >> + MAP_EVERY qid_spec_tac [‘ys’,‘xs’] >> + Induct >> + Cases_on ‘ys’ >> + rw[] >> + first_x_assum drule >> + gvs[]) >> + gvs[SUM_MAP_const]) + >- (first_x_assum irule >> + gvs[GSYM REVERSE_ZIP,MAP_REVERSE,SUM_REVERSE] >> + rename1 ‘ZIP (xs, ys)’ >> + qmatch_goalsub_abbrev_tac ‘MAP f _’ >> + irule LESS_EQ_TRANS >> + Q.SUBGOAL_THEN + ‘SUM (MAP f (ZIP (xs, ys))) + ≤ SUM(MAP (λx. cons_measure x + 1) xs) + SUM(MAP (λx. cons_measure x + 1) ys)’ + $ irule_at $ Pos hd >> + unabbrev_all_tac + >- (qpat_x_assum ‘LENGTH _ = LENGTH _’ mp_tac >> + rpt $ pop_assum kall_tac >> + MAP_EVERY qid_spec_tac [‘ys’,‘xs’] >> + Induct >> + Cases_on ‘ys’ >> + rw[] >> + first_x_assum drule >> + gvs[]) >> + gvs[SUM_MAP_const]) +QED + +Theorem eq_pure_list_alt_eq: + ∀xs. eq_pure_list xs = eq_pure_list_alt + (SUM (MAP (λ(x,y). 1 + cons_measure x + cons_measure y + + if cons_measure x < cons_measure y then 1 else 0) xs)) xs +Proof + strip_tac >> irule $ GSYM eq_pure_list_alt_thm >> rw[] +QED + +Definition cons_measure_alt_def: + cons_measure_alt x = + (case dest_Op dest_Cons x of + | NONE => 0 + | SOME (_,xs) => cons_measures_alt xs + LENGTH xs + 1) ∧ + cons_measures_alt [] = 0 ∧ + cons_measures_alt (x::xs) = cons_measure_alt x + cons_measures_alt xs +Termination + WF_REL_TAC ‘measure $ λx. sum_CASE x exp_size exp3_size’ >> + gvs[oneline clos_opTheory.dest_Op_def,AllCaseEqs(),PULL_EXISTS, + oneline clos_opTheory.dest_Cons_def,closLangTheory.exp_size_def] +End + +val pre = cv_trans_pre_rec cons_measure_alt_def + (wf_rel_tac ‘measure $ λx. sum_CASE x cv_size cv_size’ >> + cv_termination_tac >> + gvs[fetch "-" "cv_dest_Op_dest_Cons_def",AllCaseEqs(), + cvTheory.c2b_def,cvTheory.cv_lt_def0, + cvTheory.b2c_if, + oneline cvTheory.cv_ispair_def, + fetch "-" "cv_clos_op_dest_Cons_def" + ] >> + gvs[oneline cvTheory.cv_fst_def, + oneline cvTheory.cv_snd_def, + AllCaseEqs()]) + +Theorem cons_measure_alt_pre[cv_pre]: + (∀x. cons_measure_alt_pre x) ∧ + (∀v. cons_measures_alt_pre v) +Proof + ho_match_mp_tac cons_measure_alt_ind >> + rw[] >> rw[Once pre] +QED + +Triviality cons_measure_pmatch_elim = clos_opTheory.cons_measure_def + |> CONV_RULE $ QUANT_CONV $ RHS_CONV $ patternMatchesLib.PMATCH_ELIM_CONV; + +Theorem cons_measure_alt_thm: + (∀x. cons_measure_alt x = cons_measure x) ∧ + (∀v. cons_measures_alt v = SUM(MAP cons_measure v)) +Proof + reverse Induct >> rw[] >> + PURE_ONCE_REWRITE_TAC [cons_measure_pmatch_elim, cons_measure_alt_def] + >- (gvs[] >> rw[Once cons_measure_pmatch_elim]) >> + gvs[oneline clos_opTheory.dest_Op_def] >> + rpt(PURE_TOP_CASE_TAC >> gvs[]) >> + gvs[AllCaseEqs()] >> metis_tac[] +QED + +val _ = cv_auto_trans $ GSYM $ cj 1 cons_measure_alt_thm + +val _ = cv_auto_trans (eq_pure_list_alt_eq |> PURE_REWRITE_RULE [GSYM cons_measure_alt_thm]) + +val _ = cv_auto_trans clos_opTheory.int_op_def; + +Theorem cv_mul_zero: + cv_mul (Num 0) x = Num 0 +Proof + Cases_on ‘x’ \\ gvs [] +QED + +Theorem cv_lt_zero: + cv_lt x (Num 0) = Num 0 +Proof + Cases_on ‘x’ \\ gvs [] +QED + +Triviality cv_lt_eq_SUC: + cv_lt x y = Num(SUC k) ⇔ + k = 0 ∧ ∃n m. x = Num n ∧ y = Num m ∧ n < m +Proof + Cases_on ‘x’ >> gvs[cvTheory.cv_lt_def0] >> + PURE_TOP_CASE_TAC >> gvs[oneline cvTheory.b2c_def] >> + rw[] +QED + +(* why doesn't online cvTheory.cv_mul_def achieve this?*) +Theorem cv_mul_oneline: + cv_mul x y = + case (x,y) of + (Num m, Num n) => Num(m*n) + | _ => Num 0 +Proof + Cases_on ‘x’ >> Cases_on ‘y’ >> gvs[] +QED + +Theorem cv_add_oneline: + cv_add x y = + case (x,y) of + (Num m, Num n) => Num(m+n) + | (Num m, _) => Num m + | (_, Num n) => Num n + | _ => Num 0 +Proof + Cases_on ‘x’ >> Cases_on ‘y’ >> gvs[] +QED + +val pre = cv_auto_trans_pre_rec known_alt_def + (wf_rel_tac `inv_image (measure I LEX measure I) + (λx. case x of + INL (c,_,x,vs,g) => (cv_size c, cv_size x) + | INR(INL (c,_,xs,vs,g)) => (cv_size c, cv_size xs) + | INR(INR (clos,vs,c,_,g,xs)) => (cv_size c, cv_size xs))` >> + cv_termination_tac >> + gvs[cvTheory.c2b_def,oneline cvTheory.cv_ispair_def,AllCaseEqs(), + fetch "-" "cv_decide_inline_alt_def", + cvTheory.cv_eq_def, + cvTheory.cv_lt_def0, + oneline cvTheory.b2c_def + ] >> + simp[oneline cvTheory.cv_snd_def, oneline cvTheory.cv_fst_def] >> + rpt(PURE_FULL_CASE_TAC >> gvs[]) >> + gvs[cv_mul_zero, cv_lt_zero,cv_lt_eq_SUC] >> + gvs[cv_mul_oneline,AllCaseEqs(),cv_add_oneline] + >- intLib.COOPER_TAC >> + rename1 ‘m DIV 2’ >> + ‘m ≠ 0’ by(strip_tac >> gvs[]) >> + gvs[]) + +Theorem known_alt_pre[cv_pre]: + (∀inl_factor inl_max v vs g. known_alt_pre inl_factor inl_max v vs g) ∧ + (∀inl_factor inl_max v vs g. known_alts_pre inl_factor inl_max v vs g) ∧ + (∀v1 v2 v3 v4 v5 v. known_let_alts_pre v1 v2 v3 v4 v5 v) +Proof + ho_match_mp_tac known_alt_ind >> rw[] >> rw[Once pre] +QED + +Theorem known_alt_thm: + (∀inf inm v vs g c. + inf = c.inline_factor ∧ inm = c.inline_max_body_size ⇒ + known_alt c.inline_factor c.inline_max_body_size v vs g = + (λ(x,y). (HD x, y)) (known c [v] vs g)) ∧ + (∀inf inm v vs g c. + inf = c.inline_factor ∧ inm = c.inline_max_body_size ⇒ + known_alts c.inline_factor c.inline_max_body_size v vs g = + known c v vs g) ∧ + (∀clos vs inf inm g fns c. + inf = c.inline_factor ∧ inm = c.inline_max_body_size ⇒ + known_let_alts clos vs c.inline_factor c.inline_max_body_size g fns = + MAP (λ(num_args,x). + let new_vs = REPLICATE num_args Other ++ clos ++ vs in + let res = known c [x] new_vs g in + (num_args,FST (HD (FST res)))) fns) +Proof + ho_match_mp_tac known_alt_ind >> + rw[known_alt_def,clos_knownTheory.known_def] >> + rpt(pairarg_tac ORELSE IF_CASES_TAC ORELSE PURE_FULL_CASE_TAC >> gvs[decide_inline_alt_thm]) + >- (imp_res_tac clos_knownTheory.known_sing_EQ_E >> gvs[] >> + gvs[UNCURRY_eq_pair,PULL_EXISTS,PULL_FORALL] >> + gvs[clos_knownTheory.dec_inline_factor_def] >> + rpt $ first_x_assum $ qspec_then ‘c with inline_factor := c.inline_factor DIV 2’ strip_assume_tac >> gvs[]) + >- (imp_res_tac clos_knownTheory.known_sing_EQ_E >> gvs[] >> + gvs[UNCURRY_eq_pair,PULL_EXISTS,PULL_FORALL] >> + gvs[clos_knownTheory.dec_inline_factor_def] >> + rpt $ first_x_assum $ qspec_then ‘c with inline_factor := c.inline_factor DIV 2’ strip_assume_tac >> gvs[]) + >- (imp_res_tac clos_knownTheory.known_sing_EQ_E >> gvs[] >> + gvs[UNCURRY_eq_pair,PULL_EXISTS,PULL_FORALL] >> + gvs[clos_knownTheory.dec_inline_factor_def] >> + imp_res_tac clos_knownTheory.known_LENGTH_EQ_E >> + rename1 ‘known _ (_::vvs)’ >> + Cases_on ‘vvs’ >> + gvs[clos_knownTheory.known_def]) +QED + +val _ = cv_auto_trans $ + resolve_then.gen_resolve_then Any EQ_REFL (cj 2 known_alt_thm) + (fn thm => resolve_then.gen_resolve_then Any EQ_REFL thm GSYM) + +val pre = cv_auto_trans_pre clos_fvsTheory.remove_fvs_sing_def + +Theorem clos_fvs_remove_fvs_sing_pre[cv_pre]: + (∀fvs v. clos_fvs_remove_fvs_sing_pre fvs v) ∧ + (∀fvs v. clos_fvs_remove_fvs_list_pre fvs v) ∧ + (∀m fvs v. clos_fvs_remove_fvs_let_pre m fvs v) +Proof + ho_match_mp_tac clos_fvsTheory.remove_fvs_sing_ind >> rw[] >> rw[Once pre] +QED + +val _ = cv_trans $ cj 2 clos_fvsTheory.remove_fvs_sing_eq + +val _ = cv_auto_trans clos_fvsTheory.compile_def + +val pre = cv_auto_trans_pre clos_ticksTheory.remove_ticks_sing_def + +Theorem clos_ticks_remove_ticks_sing_pre[cv_pre]: + (∀v. clos_ticks_remove_ticks_sing_pre v) ∧ + (∀v. clos_ticks_remove_ticks_list_pre v) ∧ + (∀v. clos_ticks_remove_ticks_let_pre v) +Proof + ho_match_mp_tac clos_ticksTheory.remove_ticks_sing_ind >> rw[] >> rw[Once pre] +QED + +val _ = cv_trans $ cj 2 clos_ticksTheory.remove_ticks_sing_eq + +val pre = cv_auto_trans_pre clos_letopTheory.let_op_sing_def + +Theorem clos_letop_let_op_sing_pre[cv_pre]: + (∀v. clos_letop_let_op_sing_pre v) ∧ + (∀v. clos_letop_let_op_list_pre v) ∧ + (∀v. clos_letop_let_op_let_pre v) +Proof + ho_match_mp_tac clos_letopTheory.let_op_sing_ind >> + rw[] >> rw[Once pre] +QED + +val _ = cv_trans $ cj 2 clos_letopTheory.let_op_sing_eq + +val _ = cv_trans clos_knownTheory.compile_def + +(* clos_call *) + +(* TODO: free_def is duplicated between clos_known and clos_call; it + should probably be moved to a common ancestor theory *) +Theorem free_eq_free: + ∀xs. clos_known$free xs = clos_call$free xs +Proof + recInduct clos_knownTheory.free_ind >> + rw[clos_knownTheory.free_def,clos_callTheory.free_def] >> + rpt(pairarg_tac >> gvs[]) >> + rpt conj_tac + >- (ntac 2 AP_TERM_TAC >> + rw[MAP_MAP_o] >> + rw[MAP_EQ_f] >> + rpt(pairarg_tac >> gvs[]) >> + res_tac >> gvs[]) + >- (rw[MAP_MAP_o] >> + rw[MAP_EQ_f] >> + rpt(pairarg_tac >> gvs[]) >> + res_tac >> gvs[]) + >- (AP_THM_TAC >> + ntac 2 AP_TERM_TAC >> + rw[MAP_MAP_o] >> + rw[MAP_EQ_f] >> + rpt(pairarg_tac >> gvs[]) >> + res_tac >> gvs[]) +QED + +val _ = cv_trans $ GSYM free_eq_free + +val pre = cv_auto_trans_pre_rec (clos_callTheory.calls_sing_def |> PURE_REWRITE_RULE [LIST_REL_EVERY_ZIP]) + (wf_rel_tac `measure $ λx. case x of INL (e,_) => cv_size e + | INR (es,_) => cv_size es` >> + cv_termination_tac >> + irule LESS_EQ_LESS_TRANS >> + irule_at (Pos last) cv_size_map_snd >> + rw[oneline cvTheory.cv_snd_def] >> + rpt(PURE_FULL_CASE_TAC >> gvs[])) + +Theorem clos_call_calls_sing_pre[cv_pre]: + (∀v g. clos_call_calls_sing_pre v g) ∧ + (∀v g. clos_call_calls_sing_list_pre v g) +Proof + ho_match_mp_tac clos_callTheory.calls_sing_ind >> + rw[] >> rw[Once pre] >> gvs[] + >- (first_x_assum match_mp_tac >> + gvs[LIST_REL_EVERY_ZIP] >> + gvs[EVERY_MEM] >> + rw[] >> + res_tac >> + rpt(pairarg_tac >> fs[])) >> + first_x_assum match_mp_tac >> + gvs[LIST_REL_EVERY_ZIP] >> + rw[] >> res_tac >> + fs[EXISTS_MEM] >> + first_x_assum $ irule_at $ Pos hd >> + rpt(pairarg_tac >> fs[]) +QED + +val _ = cv_trans $ cj 2 clos_callTheory.calls_sing_eq + +val _ = cv_trans clos_callTheory.compile_def + +(* clos_annotate *) + +val pre = cv_auto_trans_pre clos_annotateTheory.shift_sing_def; +Theorem clos_annotate_shift_sing_pre[cv_pre]: + (∀v m l i. clos_annotate_shift_sing_pre v m l i) ∧ + (∀v m l i. clos_annotate_shift_list_pre v m l i) ∧ + (∀fns_len k new_i v. clos_annotate_shift_letrec_pre fns_len k new_i v) +Proof + ho_match_mp_tac clos_annotateTheory.shift_sing_ind + \\ rpt strip_tac \\ simp [Once pre] \\ gvs [] +QED + +val pre = cv_auto_trans_pre clos_annotateTheory.alt_free_sing_def; +Theorem clos_annotate_alt_free_sing_pre[cv_pre]: + (∀v. clos_annotate_alt_free_sing_pre v) ∧ + (∀v. clos_annotate_alt_free_list_pre v) ∧ + (∀m v. clos_annotate_alt_free_letrec_pre m v) +Proof + ho_match_mp_tac clos_annotateTheory.alt_free_sing_ind + \\ rpt strip_tac \\ simp [Once pre] \\ gvs [] +QED + +val _ = cv_trans clos_annotateTheory.annotate_sing_def; +val _ = cv_auto_trans clos_annotateTheory.compile_eq; + +(* compile_common from clos_to_bvl *) + +val _ = cv_trans clos_to_bvlTheory.chain_exps_def; +val _ = cv_auto_trans clos_to_bvlTheory.compile_common_def; + +(* bvl_jump *) + +Theorem cv_LENGTH_right_depth: + cv_LENGTH cv_xs = Num (cv_right_depth cv_xs) +Proof + gvs [cv_LENGTH_def] + \\ qsuff_tac ‘∀v n. cv_LEN v (Num n) = Num (cv_right_depth v + n)’ + \\ gvs [] \\ Induct + \\ simp [cv_right_depth_def,Once cv_LEN_def] +QED + +Theorem cv_right_depth_cv_TAKE: + ∀cv_xs n. + n ≤ cv_right_depth cv_xs ⇒ + cv_right_depth (cv_TAKE (Num n) cv_xs) = n +Proof + Induct \\ simp [cv_right_depth_def,Once cv_TAKE_def] + \\ Cases \\ gvs [cv_right_depth_def] +QED + +Theorem cv_right_depth_cv_DROP: + ∀cv_xs n. + n ≤ cv_right_depth cv_xs ⇒ + cv_right_depth (cv_DROP (Num n) cv_xs) = cv_right_depth cv_xs - n +Proof + Induct \\ simp [cv_right_depth_def,Once cv_DROP_def] + \\ Cases \\ gvs [cv_right_depth_def] +QED + +val _ = cv_auto_trans_rec bvl_jumpTheory.JumpList_def + (WF_REL_TAC ‘measure $ cv_right_depth o SND’ + \\ rpt strip_tac + \\ Cases_on ‘cv_LENGTH cv_xs’ \\ gvs [] + \\ Cases_on ‘m = 0’ \\ gvs [] + \\ Cases_on ‘m = 1’ \\ gvs [cv_LENGTH_right_depth] + \\ DEP_REWRITE_TAC [cv_right_depth_cv_TAKE,cv_right_depth_cv_DROP] + \\ gvs [DIV_LE_X,DIV_LT_X,X_LT_DIV]); + +(* clos_to_bvl *) + +val _ = cv_auto_trans clos_to_bvlTheory.init_globals_def; +val _ = cv_auto_trans clos_to_bvlTheory.init_code_def; + +Theorem cv_size_cv_map_snd_le: + ∀xs. cv_size (cv_map_snd xs) ≤ cv_size xs +Proof + Induct + \\ simp [Once cv_map_snd_def] + \\ Cases_on ‘xs’ \\ gvs [] +QED + +val pre = cv_auto_trans_pre_rec clos_to_bvlTheory.get_src_names_sing_def + (WF_REL_TAC ‘measure $ λx. case x of + | INL (e,_) => cv_size e + | INR (e,_) => cv_size e’ + \\ cv_termination_tac + \\ irule LESS_EQ_LESS_TRANS + \\ irule_at Any cv_size_cv_map_snd_le \\ gvs []); + +Theorem clos_to_bvl_get_src_names_sing_pre[cv_pre,local]: + (∀v l. clos_to_bvl_get_src_names_sing_pre v l) ∧ + (∀v l. clos_to_bvl_get_src_names_list_pre v l) +Proof + ho_match_mp_tac clos_to_bvlTheory.get_src_names_sing_ind + \\ rpt strip_tac \\ simp [Once pre] +QED + +val _ = cv_auto_trans clos_to_bvlTheory.make_name_alist_eq; + +val pre = cv_auto_trans_pre clos_to_bvlTheory.recc_Lets_def; +Theorem clos_to_bvl_recc_Lets_pre[cv_pre]: + ∀n nargs k rest. + clos_to_bvl_recc_Lets_pre n nargs k rest ⇔ k ≤ LENGTH nargs +Proof + Induct_on ‘k’ \\ gvs [] \\ simp [Once pre] \\ Cases \\ gvs [] +QED + +val _ = cv_trans clos_to_bvlTheory.recc_Let0_def; + +val pre = cv_auto_trans_pre clos_to_bvlTheory.build_recc_lets_def; +Theorem clos_to_bvl_build_recc_lets_pre[cv_pre]: + ∀nargs vs n1 fns_l c3. + clos_to_bvl_build_recc_lets_pre nargs vs n1 fns_l c3 + ⇔ + nargs ≠ [] ∧ fns_l ≤ LENGTH nargs +Proof + gvs [pre] \\ Cases using SNOC_CASES \\ gvs [REVERSE_SNOC] +QED + +val _ = cv_auto_trans clos_to_bvlTheory.add_parts_def; + +val pre = cv_auto_trans_pre_rec clos_to_bvlTheory.compile_exp_sing_def + (WF_REL_TAC ‘measure $ λx. case x of + | INL (_,e,_) => cv_size e + | INR (_,e,_) => cv_size e’ + \\ cv_termination_tac + \\ irule LESS_EQ_LESS_TRANS + \\ irule_at Any cv_size_cv_map_snd_le \\ gvs []); + +Theorem clos_to_bvl_compile_exp_sing_pre[cv_pre]: + (∀max_app v aux. clos_to_bvl_compile_exp_sing_pre max_app v aux) ∧ + (∀max_app v aux. clos_to_bvl_compile_exp_list_pre max_app v aux) +Proof + ho_match_mp_tac clos_to_bvlTheory.compile_exp_sing_ind + \\ rpt strip_tac \\ simp [Once pre] + \\ rpt strip_tac \\ gvs [cv_stdTheory.genlist_eq_GENLIST] +QED + +val _ = cv_auto_trans clos_to_bvlTheory.compile_prog_eq; + +val _ = cv_trans clos_to_bvlTheory.code_split_def; + +val _ = cv_trans_rec clos_to_bvlTheory.code_merge_def + (WF_REL_TAC ‘measure $ λ(x,y). cv_size x + cv_size y’ + \\ cv_termination_tac) ; + +Theorem cv_clos_to_bvl_code_split_imp: + ∀x y z r1 r2. + cv_clos_to_bvl_code_split x y z = Pair r1 r2 ∧ + cv_size y ≠ 0 ∧ cv_size z ≠ 0 ⇒ + cv_size r1 < cv_size x + cv_size y + cv_size z ∧ + cv_size r2 < cv_size x + cv_size y + cv_size z +Proof + Induct + \\ simp [Once $ fetch "-" "cv_clos_to_bvl_code_split_def"] + \\ rw [] \\ res_tac \\ gvs [] +QED + +val pre = cv_trans_pre_rec clos_to_bvlTheory.code_sort_def + (WF_REL_TAC ‘measure cv_size’ + \\ cv_termination_tac + \\ imp_res_tac cv_clos_to_bvl_code_split_imp \\ gvs []); + +Theorem clos_to_bvl_code_sort_pre[cv_pre]: + ∀v. clos_to_bvl_code_sort_pre v +Proof + ho_match_mp_tac clos_to_bvlTheory.code_sort_ind + \\ rpt strip_tac \\ simp [Once pre] +QED + +val _ = cv_auto_trans clos_to_bvlTheory.compile_def; + +(* to_bvl *) + +val _ = cv_auto_trans backend_asmTheory.to_bvl_def; + +(* bvl_const *) + +Definition mk_add_const_def: + mk_add_const = λx1 c2. + if c2 = 0 then x1 else Op Add [x1; Op (Const c2) []] +End + +Definition mk_add_def: + mk_add = λx1 x2. + (let + default = Op Add [x1; x2] + in + case dest_simple x2 of + NONE => + (case case_op_const x1 of + NONE => default + | SOME v5 => + case case_op_const x2 of + NONE => default + | SOME (op2,x21,n22) => + case v5 of + (op1,x11,n12) => + if op1 = Add ∧ op2 = Add then + mk_add_const (Op Add [x11; x21]) (n22 + n12) + else if op1 = Add ∧ op2 = Sub then + Op Sub + [Op Sub [x11; x21]; + Op (Const (n22 + n12)) []] + else if op1 = Sub ∧ op2 = Add then + mk_add_const (Op Sub [x11; x21]) (n22 + n12) + else default) + | SOME n2 => + case case_op_const x1 of + NONE => + (case dest_simple x1 of + NONE => mk_add_const x1 n2 + | SOME n1 => Op (Const (n2 + n1)) []) + | SOME (Add,x11,n12) => mk_add_const x11 (n2 + n12) + | SOME (Sub,x11,n12) => + Op Sub [x11; Op (Const (n2 + n12)) []] + | SOME (op,x11,n12) => default) +End + +Definition mk_sub_def: + mk_sub = λx1 x2. + (let + default = Op Sub [x1; x2] + in + case dest_simple x2 of + NONE => + (case case_op_const x1 of + NONE => default + | SOME v5 => + case case_op_const x2 of + NONE => default + | SOME (op2,x21,n22) => + case v5 of + (op1,x11,n12) => + if op1 = Add ∧ op2 = Add then + Op Add + [Op Sub [x11; x21]; + Op (Const (n22 − n12)) []] + else if op1 = Add ∧ op2 = Sub then + Op Sub + [Op Add [x11; x21]; + Op (Const (n22 − n12)) []] + else if op1 = Sub ∧ op2 = Add then + mk_add_const (Op Add [x11; x21]) (n22 − n12) + else default) + | SOME n2 => + case case_op_const x1 of + NONE => + (case dest_simple x1 of + NONE => default + | SOME n1 => Op (Const (n2 − n1)) []) + | SOME (Add,x11,n12) => + Op Sub [x11; Op (Const (n2 − n12)) []] + | SOME (Sub,x11,n12) => mk_add_const x11 (n2 − n12) + | SOME (op,x11,n12) => default) +End + +Definition mk_mul_def: + mk_mul = λx1 x2. + (let + default = Op Mult [x1; x2] + in + case dest_simple x2 of + NONE => + (case case_op_const x1 of + NONE => default + | SOME v5 => + case case_op_const x2 of + NONE => default + | SOME (op2,x21,n22) => + case v5 of + (op1,x11,n12) => + if op1 = Mult ∧ op2 = Mult then + Op Mult + [Op (Const (n22 * n12)) []; + Op Mult [x11; x21]] + else default) + | SOME n2 => + case case_op_const x1 of + NONE => + (case dest_simple x1 of + NONE => + if n2 = 1 then x1 + else if n2 = -1 then + (let + default = Op Sub [x1; Op (Const 0) []] + in + case case_op_const x1 of + NONE => + (case dest_simple x1 of + NONE => default + | SOME n1 => Op (Const (-n1)) []) + | SOME (Add,x11,n12) => + Op Sub [x11; Op (Const (-n12)) []] + | SOME (Sub,x11,n12) => + mk_add_const x11 (-n12) + | SOME (op,x11,n12) => default) + else default + | SOME n1 => Op (Const (n2 * n1)) []) + | SOME (Mult,x11,n12) => + Op Mult [x11; Op (Const (n2 * n12)) []] + | SOME (op,x11,n12) => default) +End + +Theorem SmartOp2_eq = bvl_constTheory.SmartOp2_def + |> SRULE [GSYM mk_add_const_def] + |> SRULE [Once LET_THM] + |> SRULE [GSYM mk_add_def] + |> SRULE [Once LET_THM] + |> SRULE [GSYM mk_sub_def] + |> SRULE [Once LET_THM] + |> SRULE [GSYM mk_mul_def] + |> SRULE [Once LET_THM]; + +val _ = cv_auto_trans bvl_constTheory.dest_simple_def; +val _ = cv_trans (mk_add_const_def |> SRULE [FUN_EQ_THM]); +val _ = cv_auto_trans (mk_add_def |> SRULE [FUN_EQ_THM]); +val _ = cv_trans (mk_sub_def |> SRULE [FUN_EQ_THM]); +val _ = cv_trans (mk_mul_def |> SRULE [FUN_EQ_THM]); +val _ = cv_auto_trans SmartOp2_eq; + +val pre = cv_auto_trans_pre bvl_constTheory.compile_sing_def; +Theorem bvl_const_compile_sing_pre[cv_pre]: + (∀v env. bvl_const_compile_sing_pre env v) ∧ + (∀v env. bvl_const_compile_list_pre env v) +Proof + Induct \\ rpt strip_tac \\ simp [Once pre] +QED + +val _ = cv_trans bvl_constTheory.compile_exp_eq; + +(* bvl_handle *) + +val _ = cv_trans bvl_handleTheory.SmartLet_def; +val _ = cv_auto_trans bvl_handleTheory.LetLet_def; +val _ = cv_auto_trans bvl_handleTheory.OptionalLetLet_sing_def; + +val pre = cv_auto_trans_pre bvl_handleTheory.compile_sing_def; +Theorem bvl_handle_compile_sing_pre[cv_pre,local]: + (∀v l n. bvl_handle_compile_sing_pre l n v) ∧ + (∀v l n. bvl_handle_compile_list_pre l n v) +Proof + Induct \\ rpt strip_tac \\ simp [Once pre] +QED + +val _ = cv_auto_trans bvl_handleTheory.can_raise_def; +val _ = cv_trans bvl_handleTheory.dest_handle_If_def; +val _ = cv_trans bvl_handleTheory.dest_handle_Let_def; + +val pre = cv_auto_trans_pre bvl_handleTheory.handle_adj_vars_def; +Theorem bvl_handle_handle_adj_vars_pre[cv_pre]: + (∀v l d. bvl_handle_handle_adj_vars_pre l d v) ∧ + (∀v l d. bvl_handle_handle_adj_vars1_pre l d v) +Proof + Induct \\ rpt strip_tac \\ simp [Once pre] +QED + +val pre = cv_auto_trans_pre_rec bvl_handleTheory.handle_simp_def + (WF_REL_TAC ‘measure $ λx. case x of + | INL e => cv_size e + | INR (INL e) => cv_size e + | INR (INR (e1,e2,_)) => cv_size e1 + cv_size e2 + 1’ + \\ cv_termination_tac + \\ TRY (Cases_on ‘z’ \\ gvs [] \\ NO_TAC) + \\ gvs [fetch "-" "cv_bvl_handle_dest_handle_If_def",AllCaseEqs()] + \\ gvs [fetch "-" "cv_bvl_handle_dest_handle_Let_def",AllCaseEqs()] + \\ cv_termination_tac); + +Theorem bvl_handle_handle_simp_pre[cv_pre]: + (∀v. bvl_handle_handle_simp_pre v) ∧ + (∀v. bvl_handle_handle_simp_list_pre v) ∧ + (∀x1 x2 l. bvl_handle_make_handle_pre x1 x2 l) +Proof + ho_match_mp_tac bvl_handleTheory.handle_simp_ind + \\ rpt strip_tac \\ simp [Once pre] +QED + +val _ = cv_trans bvl_handleTheory.compile_exp_eq; +val _ = bvl_handleTheory.dest_Seq_def |> CONJUNCT2 |> oneline |> cv_trans; + +val pre = cv_auto_trans_pre_rec bvl_handleTheory.compile_seqs_def + (WF_REL_TAC ‘measure $ λ(_,e,_). cv_size e’ + \\ cv_termination_tac + \\ gvs [fetch "-" "cv_bvl_handle_dest_Seq_def",AllCaseEqs()] + \\ cv_termination_tac); + +Theorem bvl_handle_compile_seqs_pre[cv_pre]: + ∀cut_size e acc. bvl_handle_compile_seqs_pre cut_size e acc +Proof + ho_match_mp_tac bvl_handleTheory.compile_seqs_ind + \\ rpt strip_tac \\ simp [Once pre] +QED + +val _ = cv_trans bvl_handleTheory.compile_any_def; + +(* bvl_inline *) + +val _ = cv_auto_trans bvl_inlineTheory.tick_inline_sing_def; + +val pre = cv_auto_trans_pre bvl_inlineTheory.is_small_sing_def; +Theorem bvl_inline_is_small_sing_pre[cv_pre,local]: + (∀v n. bvl_inline_is_small_sing_pre n v) ∧ + (∀v n. bvl_inline_is_small_list_pre n v) +Proof + Induct \\ rpt strip_tac \\ simp [Once pre] +QED + +val _ = cv_auto_trans bvl_inlineTheory.is_small_eq; + +val pre = cv_auto_trans_pre bvl_inlineTheory.is_rec_sing_def; +Theorem bvl_inline_is_rec_sing_pre[cv_pre,local]: + (∀v n. bvl_inline_is_rec_sing_pre n v) ∧ + (∀v n. bvl_inline_is_rec_list_pre n v) +Proof + Induct \\ rpt strip_tac \\ simp [Once pre] +QED + +val _ = cv_auto_trans bvl_inlineTheory.must_inline_eq; + +val pre = cv_auto_trans_pre bvl_inlineTheory.tick_inline_all_eq; +Theorem bvl_inline_tick_inline_all_pre[cv_pre,local]: + ∀v limit cs aux. bvl_inline_tick_inline_all_pre limit cs v aux +Proof + Induct \\ rpt strip_tac \\ simp [Once pre] +QED + +val pre = cv_trans bvl_inlineTheory.tick_compile_prog_def; + +val pre = cv_auto_trans_pre bvl_inlineTheory.let_op_one_def; +Theorem bvl_inline_let_op_one_pre[cv_pre,local]: + (∀v. bvl_inline_let_op_one_pre v) ∧ + (∀v. bvl_inline_let_op_list_pre v) +Proof + Induct \\ rpt strip_tac \\ simp [Once pre] +QED + +val _ = cv_trans bvl_inlineTheory.let_op_sing_eq; + +val pre = cv_auto_trans_pre bvl_inlineTheory.remove_ticks_sing_def; +Theorem bvl_inline_remove_ticks_sing_pre[cv_pre,local]: + (∀v. bvl_inline_remove_ticks_sing_pre v) ∧ + (∀v. bvl_inline_remove_ticks_list_pre v) +Proof + Induct \\ rpt strip_tac \\ simp [Once pre] +QED + +val _ = cv_trans bvl_inlineTheory.optimise_eq; +val _ = cv_auto_trans bvl_inlineTheory.compile_inc_def; +val _ = cv_trans bvl_inlineTheory.compile_prog_def; + +(* bvi_tailrec *) + +val _ = cv_trans bvi_tailrecTheory.is_rec_def; +val _ = cv_auto_trans bvi_tailrecTheory.let_wrap_def; + +Triviality term_ok_int_eq: + term_ok_int ts expr ⇔ + case expr of + | Var i => if i < LENGTH ts then EL i ts = Int else F + | Op op xs => + (if xs = [] ∧ is_const op then T else + case xs of + | [x;y] => is_arith op ∧ term_ok_int ts x ∧ term_ok_int ts y + | _ => F) + | _ => F +Proof + simp [Once bvi_tailrecTheory.term_ok_int_def] + \\ Cases_on ‘expr’ \\ gvs [] + \\ gvs [bvi_tailrecTheory.get_bin_args_def] + \\ Cases_on ‘l’ \\ gvs [] + \\ Cases_on ‘t’ \\ gvs [] + \\ Cases_on ‘t'’ \\ gvs [] +QED + +val pre = cv_auto_trans_pre term_ok_int_eq; +Theorem bvi_tailrec_term_ok_int_pre[cv_pre,local]: + ∀ts expr. bvi_tailrec_term_ok_int_pre ts expr +Proof + ho_match_mp_tac bvi_tailrecTheory.term_ok_int_ind + \\ rw [] \\ simp [Once pre] \\ rw [] + \\ gvs [bvi_tailrecTheory.get_bin_args_def] +QED + +val _ = cv_trans bvi_tailrecTheory.get_bin_args_def; + +Triviality cv_bvi_tailrec_get_bin_args_lemma: + cv_bvi_tailrec_get_bin_args (Pair t v) = Pair x (Pair y z) ⇒ + cv_size y ≤ cv_size v ∧ + cv_size z ≤ cv_size v +Proof + gvs [fetch "-" "cv_bvi_tailrec_get_bin_args_def",AllCaseEqs()] + \\ cv_termination_tac +QED + +val pre = cv_auto_trans_pre_rec bvi_tailrecTheory.term_ok_any_def + (WF_REL_TAC ‘measure $ λ(_,_,e). cv_size e’ + \\ cv_termination_tac + \\ Cases_on ‘z’ \\ gvs [] + \\ imp_res_tac cv_bvi_tailrec_get_bin_args_lemma + \\ gvs []); + +Theorem bvi_tailrec_term_ok_any_pre[cv_pre]: + ∀ts list v. bvi_tailrec_term_ok_any_pre ts list v +Proof + ho_match_mp_tac bvi_tailrecTheory.term_ok_any_ind + \\ rpt strip_tac \\ simp [Once pre] + \\ rpt strip_tac \\ gvs [] +QED + +val pre = cv_auto_trans_pre bvi_tailrecTheory.scan_expr_sing_def; +Theorem bvi_tailrec_scan_expr_sing_pre[cv_pre]: + (∀v ts loc. bvi_tailrec_scan_expr_sing_pre ts loc v) ∧ + (∀v ts loc. bvi_tailrec_scan_expr_list_pre ts loc v) +Proof + Induct \\ rpt strip_tac \\ simp [Once pre] +QED + +val pre = cv_auto_trans_pre bvi_tailrecTheory.rewrite_eq; +Theorem bvi_tailrec_rewrite_pre[cv_pre]: + ∀loc next opr acc ts v. bvi_tailrec_rewrite_pre loc next opr acc ts v +Proof + ho_match_mp_tac bvi_tailrecTheory.rewrite_ind + \\ rpt strip_tac \\ simp [Once pre] + \\ rpt strip_tac \\ gvs [bvi_tailrecTheory.scan_expr_eq] +QED + +val pre = cv_auto_trans_pre bvi_tailrecTheory.has_rec_sing_def; +Theorem bvi_tailrec_has_rec_sing_pre[cv_pre]: + (∀v loc. bvi_tailrec_has_rec_sing_pre loc v) ∧ + (∀v loc. bvi_tailrec_has_rec_list_pre loc v) +Proof + Induct \\ rpt strip_tac \\ simp [Once pre] +QED + +val _ = cv_auto_trans bvi_tailrecTheory.check_exp_eq; +val _ = cv_auto_trans bvi_tailrecTheory.compile_exp_def; + +val pre = cv_auto_trans_pre bvi_tailrecTheory.compile_prog_def; +Theorem bvi_tailrec_compile_prog_pre[cv_pre]: + ∀next v. bvi_tailrec_compile_prog_pre next v +Proof + ho_match_mp_tac bvi_tailrecTheory.compile_prog_ind + \\ rpt strip_tac \\ simp [Once pre] +QED + +(* bvi_let *) + +Theorem cv_size_LAST: + ∀g. cv_size (cv_LAST g) ≤ cv_size g +Proof + Induct \\ simp [Once cv_LAST_def] \\ rw [] +QED + +Theorem cv_size_FRONT: + ∀g. cv_size (cv_FRONT g) ≤ cv_size g +Proof + Induct \\ simp [Once cv_FRONT_def] \\ rw [] +QED + +val pre = cv_auto_trans_pre_rec bvi_letTheory.compile_sing_def + (WF_REL_TAC ‘measure $ λx. case x of + | INL (_,_,e) => cv_size e + | INR (_,_,e) => cv_size e’ + \\ cv_termination_tac + \\ gvs [cv_LENGTH_right_depth] + \\ Cases_on ‘z’ \\ gvs [cv_right_depth_def] + \\ irule LESS_EQ_LESS_TRANS + \\ (irule_at Any cv_size_FRONT ORELSE + irule_at Any cv_size_LAST) + \\ gvs []); + +Theorem bvi_let_compile_sing_pre[cv_pre]: + (∀env d v. bvi_let_compile_sing_pre env d v) ∧ + (∀env d v. bvi_let_compile_list_pre env d v) +Proof + ho_match_mp_tac bvi_letTheory.compile_sing_ind + \\ rpt strip_tac \\ simp [Once pre] +QED + +val _ = cv_trans bvi_letTheory.compile_exp_eq; + +(* bvl_to_bvi *) + +val _ = cv_auto_trans bvl_to_bviTheory.get_names_def; +val _ = cv_auto_trans bvl_to_bviTheory.stubs_def; +val _ = cv_trans bvl_to_bviTheory.destLet_def; + +val cv_bvl_to_bvi_destLet_def = fetch "-" "cv_bvl_to_bvi_destLet_def"; + +Theorem destLet_lemma[local]: + cv_bvl_to_bvi_destLet (cv_fst z) = Pair x1 x2 ⇒ + cv_size x1 ≤ cv_size z ∧ + cv_size x2 ≤ cv_size z +Proof + Cases_on ‘z’ \\ gvs [cv_bvl_to_bvi_destLet_def] + \\ rename [‘_ ≤ cv_size g1 + (cv_size g2 + 1)’] + \\ Cases_on ‘g1’ \\ gvs [] + \\ Cases_on ‘g’ \\ gvs [AllCaseEqs()] \\ rw [] + \\ Cases_on ‘g'’ \\ gvs [] +QED + +Theorem destLet_lemma2[local]: + cv_bvl_to_bvi_destLet (cv_snd z) = Pair x1 x2 ⇒ + cv_size x1 ≤ cv_size z ∧ + cv_size x2 ≤ cv_size z +Proof + Cases_on ‘z’ \\ gvs [cv_bvl_to_bvi_destLet_def] + \\ rename [‘_ ≤ cv_size g1 + (cv_size g2 + 1)’] + \\ Cases_on ‘g2’ \\ gvs [] + \\ Cases_on ‘g’ \\ gvs [AllCaseEqs()] \\ rw [] + \\ Cases_on ‘g'’ \\ gvs [] +QED + +val pre = cv_auto_trans_pre_rec bvl_to_bviTheory.compile_exps_sing_def + (WF_REL_TAC ‘measure $ λx. case x of + | INL (_,e) => cv_size e + | INR (_,e) => cv_size e’ + \\ cv_termination_tac + \\ imp_res_tac destLet_lemma + \\ imp_res_tac destLet_lemma2 + \\ gvs []); + +Theorem bvl_to_bvi_compile_exps_sing_pre[cv_pre]: + (∀n v. bvl_to_bvi_compile_exps_sing_pre n v) ∧ + (∀n v. bvl_to_bvi_compile_exps_list_pre n v) +Proof + ho_match_mp_tac bvl_to_bviTheory.compile_exps_sing_ind + \\ rpt strip_tac \\ simp [Once pre] +QED + +val _ = cv_trans bvl_to_bviTheory.compile_single_eq; + +val pre = cv_trans_pre bvl_to_bviTheory.global_count_sing_def; +Theorem bvl_to_bvi_global_count_sing_pre[cv_pre]: + (∀v. bvl_to_bvi_global_count_sing_pre v) ∧ + (∀v. bvl_to_bvi_global_count_list_pre v) +Proof + Induct \\ rpt strip_tac \\ simp [Once pre] +QED + +val _ = cv_auto_trans bvl_to_bviTheory.compile_prog_eq; +val _ = cv_auto_trans bvl_to_bviTheory.compile_def; + +(* to_bvi *) + +val _ = cv_auto_trans backend_asmTheory.to_bvi_def; + +(* to_data *) + +Definition num_size_acc_def: + num_size_acc n acc = + if n < 2 ** 32 then acc + 2:num else num_size_acc (n DIV 2 ** 32) (acc + 1) +End + +Theorem num_size_eq_acc: + num_size n = num_size_acc n 0 +Proof + qsuff_tac ‘∀n acc. num_size_acc n acc = num_size n + acc’ \\ gvs [] + \\ ho_match_mp_tac data_spaceTheory.num_size_ind \\ rw [] + \\ once_rewrite_tac [data_spaceTheory.num_size_def, num_size_acc_def] + \\ rw [] \\ gvs [] +QED + +val _ = cv_trans num_size_acc_def; +val _ = cv_trans num_size_eq_acc; + +val pre = cv_auto_trans_pre data_liveTheory.compile_def; +Theorem data_live_compile_pre[cv_pre,local]: + ∀v live. data_live_compile_pre v live +Proof + ho_match_mp_tac data_liveTheory.compile_ind \\ rw [] \\ simp [Once pre] +QED + +val _ = cv_auto_trans bvi_to_dataTheory.optimise_def; +val _ = cv_auto_trans bvi_to_dataTheory.op_requires_names_eqn; + +val pre = cv_auto_trans_pre bvi_to_dataTheory.compile_sing_def; +Theorem bvi_to_data_compile_sing_pre[cv_pre,local]: + (∀n env tail live v. bvi_to_data_compile_sing_pre n env tail live v) ∧ + (∀n env live v. bvi_to_data_compile_list_pre n env live v) Proof - cheat + ho_match_mp_tac bvi_to_dataTheory.compile_sing_ind + \\ rpt strip_tac \\ simp [Once pre] QED -val _ = cv_trans to_data_fake; +val _ = cv_auto_trans rich_listTheory.COUNT_LIST_GENLIST; +val _ = cv_trans bvi_to_dataTheory.compile_exp_eq; +val _ = cv_auto_trans bvi_to_dataTheory.compile_prog_def; +val _ = cv_trans to_data_def; val _ = Feedback.set_trace "TheoryPP.include_docs" 0; val _ = export_theory(); diff --git a/examples/compilation/ag32/Holmakefile b/examples/compilation/ag32/Holmakefile index 62a8fe122c..348ae8f9fd 100644 --- a/examples/compilation/ag32/Holmakefile +++ b/examples/compilation/ag32/Holmakefile @@ -1,4 +1,4 @@ -INCLUDES = $(CAKEMLDIR)/misc $(CAKEMLDIR)/basis $(CAKEMLDIR)/compiler ../.. +INCLUDES = $(CAKEMLDIR)/misc $(CAKEMLDIR)/basis $(CAKEMLDIR)/compiler $(CAKEMLDIR)/cv_translator ../.. CLINE_OPTIONS = -j1 all: $(DEFAULT_TARGETS) README.md @@ -8,3 +8,7 @@ README_SOURCES = $(wildcard *Script.sml) $(wildcard *Lib.sml) $(wildcard *Syntax DIRS = $(wildcard */) README.md: $(CAKEMLDIR)/developers/readme_gen readmePrefix $(patsubst %,%readmePrefix,$(DIRS)) $(README_SOURCES) $(protect $(CAKEMLDIR)/developers/readme_gen) $(README_SOURCES) + +ifdef POLY +HOLHEAP = $(CAKEMLDIR)/cv_translator/cake_compile_heap +endif diff --git a/examples/compilation/ag32/helloCompileScript.sml b/examples/compilation/ag32/helloCompileScript.sml index c0a0929aa5..8ef93796b5 100644 --- a/examples/compilation/ag32/helloCompileScript.sml +++ b/examples/compilation/ag32/helloCompileScript.sml @@ -1,11 +1,11 @@ (* Compiles the hello example by evaluation inside the logic of HOL *) -open preamble compilationLib helloProgTheory +open preamble helloProgTheory eval_cake_compile_ag32Lib val _ = new_theory "helloCompile" -val hello_compiled = save_thm("hello_compiled", - compile_ag32 "hello" hello_prog_def); +Theorem hello_compiled = + eval_cake_compile_ag32 "" hello_prog_def "hello.S"; val _ = export_theory (); diff --git a/examples/compilation/ag32/proofs/helloProofScript.sml b/examples/compilation/ag32/proofs/helloProofScript.sml index 5aafad6b5b..e66c686ace 100644 --- a/examples/compilation/ag32/proofs/helloProofScript.sml +++ b/examples/compilation/ag32/proofs/helloProofScript.sml @@ -20,30 +20,37 @@ val hello_io_events_def = val (hello_sem,hello_output) = hello_io_events_def |> SPEC_ALL |> UNDISCH |> CONJ_PAIR val (hello_not_fail,hello_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail hello_sem |> CONJ_PAIR -val ffi_names = - ``config.lab_conf.ffi_names`` - |> (REWRITE_CONV[helloCompileTheory.config_def] THENC EVAL) +val ffinames_to_string_list_def = backendTheory.ffinames_to_string_list_def; -val LENGTH_code = - ``LENGTH code`` - |> (REWRITE_CONV[helloCompileTheory.code_def] THENC listLib.LENGTH_CONV) +Theorem extcalls_ffi_names: + extcalls info.lab_conf.ffi_names = ffis +Proof + rewrite_tac [hello_compiled] + \\ qspec_tac (‘info.lab_conf.ffi_names’,‘xs’) \\ Cases + \\ gvs [extcalls_def,ffinames_to_string_list_def,libTheory.the_def] + \\ Induct_on ‘x’ + \\ gvs [extcalls_def,ffinames_to_string_list_def,libTheory.the_def] + \\ Cases \\ gvs [extcalls_def,ffinames_to_string_list_def,libTheory.the_def] +QED -val LENGTH_data = - ``LENGTH data`` - |> (REWRITE_CONV[helloCompileTheory.data_def] THENC listLib.LENGTH_CONV) +val ffis = ffis_def |> CONV_RULE (RAND_CONV EVAL); +val ffi_names = extcalls_ffi_names |> SRULE [ffis] -val shmem = - ``config.lab_conf.shmem_extra`` - |> (REWRITE_CONV[helloCompileTheory.config_def] THENC EVAL) +val LENGTH_code = “LENGTH code” |> SCONV [hello_compiled]; +val LENGTH_data = “LENGTH data” |> SCONV [hello_compiled]; +val shmem = “info.lab_conf.shmem_extra” |> SCONV [hello_compiled]; Overload hello_machine_config = - ``ag32_machine_config (extcalls config.lab_conf.ffi_names) (LENGTH code) (LENGTH data)`` + “ag32_machine_config (extcalls info.lab_conf.ffi_names) (LENGTH code) (LENGTH data)” Theorem target_state_rel_hello_start_asm_state: SUM (MAP strlen cl) + LENGTH cl ≤ cline_size ∧ LENGTH inp ≤ stdin_size ∧ - is_ag32_init_state (init_memory code data (extcalls config.lab_conf.ffi_names) (cl,inp)) ms ⇒ - ∃n. target_state_rel ag32_target (init_asm_state code data (extcalls config.lab_conf.ffi_names) (cl,inp)) (FUNPOW Next n ms) ∧ + is_ag32_init_state (init_memory code data (extcalls info.lab_conf.ffi_names) + (cl,inp)) ms ⇒ + ∃n. target_state_rel ag32_target + (init_asm_state code data (extcalls info.lab_conf.ffi_names) (cl,inp)) + (FUNPOW Next n ms) ∧ ((FUNPOW Next n ms).io_events = ms.io_events) ∧ (∀x. x ∉ (ag32_startup_addresses) ⇒ ((FUNPOW Next n ms).MEM x = ms.MEM x)) @@ -52,7 +59,7 @@ Proof \\ drule (GEN_ALL init_asm_state_RTC_asm_step) \\ disch_then drule \\ simp_tac std_ss [] - \\ disch_then(qspecl_then[`code`,`data`,`extcalls config.lab_conf.ffi_names`]mp_tac) + \\ disch_then(qspecl_then[`code`,`data`,`extcalls info.lab_conf.ffi_names`]mp_tac) \\ impl_tac >- ( EVAL_TAC>> fs[ffi_names,LENGTH_data,LENGTH_code,extcalls_def,shmem]) \\ strip_tac \\ drule (GEN_ALL target_state_rel_ag32_init) @@ -69,8 +76,9 @@ val hello_startup_clock_def = |> SIMP_RULE bool_ss [GSYM RIGHT_EXISTS_IMP_THM,SKOLEM_THM]); val compile_correct_applied = - MATCH_MP compile_correct hello_compiled - |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] + MATCH_MP compile_correct (cj 1 hello_compiled) + |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm, + GSYM AND_IMP_INTRO] |> C MATCH_MP hello_not_fail |> C MATCH_MP ag32_backend_config_ok |> REWRITE_RULE[hello_sem_sing,AND_IMP_INTRO] @@ -81,29 +89,33 @@ val compile_correct_applied = |> Q.GEN`cbspace` |> Q.SPEC`0` |> Q.GEN`data_sp` |> Q.SPEC`0` -Triviality to_MAP_ExtCall: - [ExtCall n] = MAP ExtCall [n] ∧ - (ExtCall n::MAP ExtCall ns) = MAP ExtCall (n::ns) -Proof - fs [] -QED - Theorem hello_installed: SUM (MAP strlen cl) + LENGTH cl ≤ cline_size ∧ LENGTH inp ≤ stdin_size ∧ - is_ag32_init_state (init_memory code data (extcalls config.lab_conf.ffi_names) (cl,inp)) ms0 ⇒ - installed code 0 data 0 config.lab_conf.ffi_names + is_ag32_init_state (init_memory code data (extcalls info.lab_conf.ffi_names) (cl,inp)) ms0 ⇒ + installed code 0 data 0 info.lab_conf.ffi_names (heap_regs ag32_backend_config.stack_conf.reg_names) - (hello_machine_config) config.lab_conf.shmem_extra + (hello_machine_config) info.lab_conf.shmem_extra (FUNPOW Next (hello_startup_clock ms0 inp cl) ms0) Proof rewrite_tac[ffi_names, extcalls_def, shmem] \\ strip_tac - \\ rewrite_tac [to_MAP_ExtCall] + \\ qmatch_asmsub_abbrev_tac ‘init_memory _ _ ff’ + \\ qmatch_goalsub_abbrev_tac ‘installed _ _ _ _ dd’ + \\ ‘dd = SOME (MAP ExtCall ff)’ by + (unabbrev_all_tac + \\ assume_tac (cj 1 hello_compiled) + \\ drule ag32_configProofTheory.compile_imp_ffi_names + \\ gvs [hello_compiled] + \\ gvs [GSYM hello_compiled,ffis] + \\ simp [backendTheory.set_oracle_def, + ag32_configTheory.ag32_backend_config_def]) + \\ asm_rewrite_tac [] \\ irule ag32_installed \\ drule hello_startup_clock_def \\ disch_then drule \\ rewrite_tac[ffi_names, extcalls_def] + \\ unabbrev_all_tac \\ disch_then drule \\ strip_tac \\ simp[] @@ -161,7 +173,8 @@ QED Theorem hello_ag32_next: SUM (MAP strlen cl) + LENGTH cl ≤ cline_size ∧ wfcl cl ∧ LENGTH inp ≤ stdin_size ∧ - is_ag32_init_state (init_memory code data (extcalls config.lab_conf.ffi_names) (cl,inp)) ms0 + is_ag32_init_state (init_memory code data (extcalls info.lab_conf.ffi_names) + (cl,inp)) ms0 ⇒ ∃k1. ∀k. k1 ≤ k ⇒ let ms = FUNPOW Next k ms0 in diff --git a/examples/compilation/ag32/proofs/sortProofScript.sml b/examples/compilation/ag32/proofs/sortProofScript.sml index 3ac0a6c226..276bfd24dd 100644 --- a/examples/compilation/ag32/proofs/sortProofScript.sml +++ b/examples/compilation/ag32/proofs/sortProofScript.sml @@ -37,30 +37,34 @@ val sort_io_events_def = val (sort_sem,sort_output) = sort_io_events_def |> SPEC_ALL |> CONJ_PAIR val (sort_not_fail,sort_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail sort_sem |> CONJ_PAIR -val ffi_names = - ``config.lab_conf.ffi_names`` - |> (REWRITE_CONV[sortCompileTheory.config_def] THENC EVAL) +val ffinames_to_string_list_def = backendTheory.ffinames_to_string_list_def; -val LENGTH_code = - ``LENGTH code`` - |> (REWRITE_CONV[sortCompileTheory.code_def] THENC listLib.LENGTH_CONV) +Theorem extcalls_ffi_names: + extcalls info.lab_conf.ffi_names = ffis +Proof + rewrite_tac [sort_compiled] + \\ qspec_tac (‘info.lab_conf.ffi_names’,‘xs’) \\ Cases + \\ gvs [extcalls_def,ffinames_to_string_list_def,libTheory.the_def] + \\ Induct_on ‘x’ + \\ gvs [extcalls_def,ffinames_to_string_list_def,libTheory.the_def] + \\ Cases \\ gvs [extcalls_def,ffinames_to_string_list_def,libTheory.the_def] +QED -val LENGTH_data = - ``LENGTH data`` - |> (REWRITE_CONV[sortCompileTheory.data_def] THENC listLib.LENGTH_CONV) +val ffis = ffis_def |> CONV_RULE (RAND_CONV EVAL); +val ffi_names = extcalls_ffi_names |> SRULE [ffis] -val shmem = - ``config.lab_conf.shmem_extra`` - |> (REWRITE_CONV[sortCompileTheory.config_def] THENC EVAL) +val LENGTH_code = “LENGTH code” |> SCONV [sort_compiled]; +val LENGTH_data = “LENGTH data” |> SCONV [sort_compiled]; +val shmem = “info.lab_conf.shmem_extra” |> SCONV [sort_compiled]; Overload sort_machine_config = - ``ag32_machine_config (extcalls config.lab_conf.ffi_names) (LENGTH code) (LENGTH data)`` + “ag32_machine_config (extcalls info.lab_conf.ffi_names) (LENGTH code) (LENGTH data)” Theorem target_state_rel_sort_start_asm_state: SUM (MAP strlen cl) + LENGTH cl ≤ cline_size ∧ LENGTH inp ≤ stdin_size ∧ - is_ag32_init_state (init_memory code data (extcalls config.lab_conf.ffi_names) (cl,inp)) ms ⇒ - ∃n. target_state_rel ag32_target (init_asm_state code data (extcalls config.lab_conf.ffi_names) (cl,inp)) (FUNPOW Next n ms) ∧ + is_ag32_init_state (init_memory code data (extcalls info.lab_conf.ffi_names) (cl,inp)) ms ⇒ + ∃n. target_state_rel ag32_target (init_asm_state code data (extcalls info.lab_conf.ffi_names) (cl,inp)) (FUNPOW Next n ms) ∧ ((FUNPOW Next n ms).io_events = ms.io_events) ∧ (∀x. x ∉ (ag32_startup_addresses) ⇒ ((FUNPOW Next n ms).MEM x = ms.MEM x)) @@ -69,7 +73,7 @@ Proof \\ drule (GEN_ALL init_asm_state_RTC_asm_step) \\ disch_then drule \\ simp_tac std_ss [] - \\ disch_then(qspecl_then[`code`,`data`,`extcalls config.lab_conf.ffi_names`]mp_tac) + \\ disch_then(qspecl_then[`code`,`data`,`extcalls info.lab_conf.ffi_names`]mp_tac) \\ impl_tac >- ( EVAL_TAC>> fs[ffi_names,LENGTH_data,LENGTH_code,extcalls_def]) \\ strip_tac \\ drule (GEN_ALL target_state_rel_ag32_init) @@ -86,8 +90,9 @@ val sort_startup_clock_def = |> SIMP_RULE bool_ss [GSYM RIGHT_EXISTS_IMP_THM,SKOLEM_THM]); val compile_correct_applied = - MATCH_MP compile_correct sort_compiled - |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] + MATCH_MP compile_correct (cj 1 sort_compiled) + |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm, + GSYM AND_IMP_INTRO] |> C MATCH_MP sort_not_fail |> C MATCH_MP ag32_backend_config_ok |> REWRITE_RULE[sort_sem_sing,AND_IMP_INTRO] @@ -98,29 +103,33 @@ val compile_correct_applied = |> Q.GEN`cbspace` |> Q.SPEC`0` |> Q.GEN`data_sp` |> Q.SPEC`0` -Triviality to_MAP_ExtCall: - [ExtCall n] = MAP ExtCall [n] ∧ - (ExtCall n::MAP ExtCall ns) = MAP ExtCall (n::ns) -Proof - fs [] -QED - Theorem sort_installed: SUM (MAP strlen cl) + LENGTH cl ≤ cline_size ∧ LENGTH inp ≤ stdin_size ∧ - is_ag32_init_state (init_memory code data (extcalls config.lab_conf.ffi_names) (cl,inp)) ms0 ⇒ - installed code 0 data 0 config.lab_conf.ffi_names + is_ag32_init_state (init_memory code data (extcalls info.lab_conf.ffi_names) (cl,inp)) ms0 ⇒ + installed code 0 data 0 info.lab_conf.ffi_names (heap_regs ag32_backend_config.stack_conf.reg_names) - (sort_machine_config) config.lab_conf.shmem_extra + (sort_machine_config) info.lab_conf.shmem_extra (FUNPOW Next (sort_startup_clock ms0 inp cl) ms0) Proof rewrite_tac[ffi_names, extcalls_def, shmem] \\ strip_tac - \\ rewrite_tac [to_MAP_ExtCall] + \\ qmatch_asmsub_abbrev_tac ‘init_memory _ _ ff’ + \\ qmatch_goalsub_abbrev_tac ‘installed _ _ _ _ dd’ + \\ ‘dd = SOME (MAP ExtCall ff)’ by + (unabbrev_all_tac + \\ assume_tac (cj 1 sort_compiled) + \\ drule ag32_configProofTheory.compile_imp_ffi_names + \\ gvs [sort_compiled] + \\ gvs [GSYM sort_compiled,ffis] + \\ simp [backendTheory.set_oracle_def, + ag32_configTheory.ag32_backend_config_def]) + \\ asm_rewrite_tac [] \\ irule ag32_installed \\ drule sort_startup_clock_def \\ disch_then drule \\ rewrite_tac[ffi_names, extcalls_def] + \\ unabbrev_all_tac \\ disch_then drule \\ strip_tac \\ simp[] @@ -186,7 +195,7 @@ QED Theorem sort_ag32_next: LENGTH inp ≤ stdin_size ∧ - is_ag32_init_state (init_memory code data (extcalls config.lab_conf.ffi_names) ([strlit"sort"],inp)) ms0 + is_ag32_init_state (init_memory code data (extcalls info.lab_conf.ffi_names) ([strlit"sort"],inp)) ms0 ⇒ ∃k1. ∀k. k1 ≤ k ⇒ let ms = FUNPOW Next k ms0 in diff --git a/examples/compilation/ag32/proofs/wordcountProofScript.sml b/examples/compilation/ag32/proofs/wordcountProofScript.sml index ba8943f934..011e07c99a 100644 --- a/examples/compilation/ag32/proofs/wordcountProofScript.sml +++ b/examples/compilation/ag32/proofs/wordcountProofScript.sml @@ -47,30 +47,34 @@ val wordcount_io_events_def = val (wordcount_sem,wordcount_output) = wordcount_io_events_def |> SPEC_ALL |> CONJ_PAIR val (wordcount_not_fail,wordcount_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail wordcount_sem |> CONJ_PAIR -val ffi_names = - ``config.lab_conf.ffi_names`` - |> (REWRITE_CONV[wordcountCompileTheory.config_def] THENC EVAL) +val ffinames_to_string_list_def = backendTheory.ffinames_to_string_list_def; -val LENGTH_code = - ``LENGTH code`` - |> (REWRITE_CONV[wordcountCompileTheory.code_def] THENC listLib.LENGTH_CONV) +Theorem extcalls_ffi_names: + extcalls info.lab_conf.ffi_names = ffis +Proof + rewrite_tac [wordcount_compiled] + \\ qspec_tac (‘info.lab_conf.ffi_names’,‘xs’) \\ Cases + \\ gvs [extcalls_def,ffinames_to_string_list_def,libTheory.the_def] + \\ Induct_on ‘x’ + \\ gvs [extcalls_def,ffinames_to_string_list_def,libTheory.the_def] + \\ Cases \\ gvs [extcalls_def,ffinames_to_string_list_def,libTheory.the_def] +QED -val LENGTH_data = - ``LENGTH data`` - |> (REWRITE_CONV[wordcountCompileTheory.data_def] THENC listLib.LENGTH_CONV) +val ffis = ffis_def |> CONV_RULE (RAND_CONV EVAL); +val ffi_names = extcalls_ffi_names |> SRULE [ffis] -val shmem = - ``config.lab_conf.shmem_extra`` - |> (REWRITE_CONV[wordcountCompileTheory.config_def] THENC EVAL) +val LENGTH_code = “LENGTH code” |> SCONV [wordcount_compiled]; +val LENGTH_data = “LENGTH data” |> SCONV [wordcount_compiled]; +val shmem = “info.lab_conf.shmem_extra” |> SCONV [wordcount_compiled]; Overload wordcount_machine_config = - ``ag32_machine_config (extcalls config.lab_conf.ffi_names) (LENGTH code) (LENGTH data)`` + “ag32_machine_config (extcalls info.lab_conf.ffi_names) (LENGTH code) (LENGTH data)” Theorem target_state_rel_wordcount_start_asm_state: SUM (MAP strlen cl) + LENGTH cl ≤ cline_size ∧ LENGTH inp ≤ stdin_size ∧ - is_ag32_init_state (init_memory code data (extcalls config.lab_conf.ffi_names) (cl,inp)) ms ⇒ - ∃n. target_state_rel ag32_target (init_asm_state code data (extcalls config.lab_conf.ffi_names) (cl,inp)) (FUNPOW Next n ms) ∧ + is_ag32_init_state (init_memory code data (extcalls info.lab_conf.ffi_names) (cl,inp)) ms ⇒ + ∃n. target_state_rel ag32_target (init_asm_state code data (extcalls info.lab_conf.ffi_names) (cl,inp)) (FUNPOW Next n ms) ∧ ((FUNPOW Next n ms).io_events = ms.io_events) ∧ (∀x. x ∉ (ag32_startup_addresses) ⇒ ((FUNPOW Next n ms).MEM x = ms.MEM x)) @@ -79,7 +83,7 @@ Proof \\ drule (GEN_ALL init_asm_state_RTC_asm_step) \\ disch_then drule \\ simp_tac std_ss [] - \\ disch_then(qspecl_then[`code`,`data`,`extcalls config.lab_conf.ffi_names`]mp_tac) + \\ disch_then(qspecl_then[`code`,`data`,`extcalls info.lab_conf.ffi_names`]mp_tac) \\ impl_tac >- ( EVAL_TAC>> fs[ffi_names,LENGTH_data,LENGTH_code,extcalls_def]) \\ strip_tac \\ drule (GEN_ALL target_state_rel_ag32_init) @@ -96,8 +100,9 @@ val wordcount_startup_clock_def = |> SIMP_RULE bool_ss [GSYM RIGHT_EXISTS_IMP_THM,SKOLEM_THM]); val wordcount_compile_correct_applied = - MATCH_MP compile_correct wordcount_compiled - |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] + MATCH_MP compile_correct (cj 1 wordcount_compiled) + |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm, + GSYM AND_IMP_INTRO] |> C MATCH_MP wordcount_not_fail |> C MATCH_MP ag32_backend_config_ok |> REWRITE_RULE[wordcount_sem_sing,AND_IMP_INTRO] @@ -109,29 +114,33 @@ val wordcount_compile_correct_applied = |> Q.GEN`data_sp` |> Q.SPEC`0` |> curry save_thm "wordcount_compile_correct_applied"; -Triviality to_MAP_ExtCall: - [ExtCall n] = MAP ExtCall [n] ∧ - (ExtCall n::MAP ExtCall ns) = MAP ExtCall (n::ns) -Proof - fs [] -QED - Theorem wordcount_installed: SUM (MAP strlen cl) + LENGTH cl ≤ cline_size ∧ LENGTH inp ≤ stdin_size ∧ - is_ag32_init_state (init_memory code data (extcalls config.lab_conf.ffi_names) (cl,inp)) ms0 ⇒ - installed code 0 data 0 config.lab_conf.ffi_names + is_ag32_init_state (init_memory code data (extcalls info.lab_conf.ffi_names) (cl,inp)) ms0 ⇒ + installed code 0 data 0 info.lab_conf.ffi_names (heap_regs ag32_backend_config.stack_conf.reg_names) - (wordcount_machine_config) config.lab_conf.shmem_extra + (wordcount_machine_config) info.lab_conf.shmem_extra (FUNPOW Next (wordcount_startup_clock ms0 inp cl) ms0) Proof rewrite_tac[ffi_names, extcalls_def, shmem] \\ strip_tac - \\ rewrite_tac [to_MAP_ExtCall] + \\ qmatch_asmsub_abbrev_tac ‘init_memory _ _ ff’ + \\ qmatch_goalsub_abbrev_tac ‘installed _ _ _ _ dd’ + \\ ‘dd = SOME (MAP ExtCall ff)’ by + (unabbrev_all_tac + \\ assume_tac (cj 1 wordcount_compiled) + \\ drule ag32_configProofTheory.compile_imp_ffi_names + \\ gvs [wordcount_compiled] + \\ gvs [GSYM wordcount_compiled,ffis] + \\ simp [backendTheory.set_oracle_def, + ag32_configTheory.ag32_backend_config_def]) + \\ asm_rewrite_tac [] \\ irule ag32_installed \\ drule wordcount_startup_clock_def \\ disch_then drule \\ rewrite_tac[ffi_names, extcalls_def] + \\ unabbrev_all_tac \\ disch_then drule \\ strip_tac \\ simp[] @@ -195,7 +204,7 @@ QED Theorem wordcount_ag32_next: LENGTH inp ≤ stdin_size ∧ - is_ag32_init_state (init_memory code data (extcalls config.lab_conf.ffi_names) ([strlit"wordcount"],inp)) ms0 + is_ag32_init_state (init_memory code data (extcalls info.lab_conf.ffi_names) ([strlit"wordcount"],inp)) ms0 ⇒ ∃k1. ∀k. k1 ≤ k ⇒ let ms = FUNPOW Next k ms0 in diff --git a/examples/compilation/ag32/sortCompileScript.sml b/examples/compilation/ag32/sortCompileScript.sml index 635af1728e..b429e60897 100644 --- a/examples/compilation/ag32/sortCompileScript.sml +++ b/examples/compilation/ag32/sortCompileScript.sml @@ -1,11 +1,11 @@ (* Compiles the sort example by evaluation inside the logic of HOL *) -open preamble compilationLib sortProgTheory +open preamble sortProgTheory eval_cake_compile_ag32Lib val _ = new_theory "sortCompile" -val sort_compiled = save_thm("sort_compiled", - compile_ag32 "sort" sort_prog_def); +Theorem sort_compiled = + eval_cake_compile_ag32 "" sort_prog_def "sort.S"; val _ = export_theory (); diff --git a/examples/compilation/ag32/wordcountCompileScript.sml b/examples/compilation/ag32/wordcountCompileScript.sml index dfc16e5e54..d630ba8303 100644 --- a/examples/compilation/ag32/wordcountCompileScript.sml +++ b/examples/compilation/ag32/wordcountCompileScript.sml @@ -2,12 +2,11 @@ Compile the wordcount program to machine code by evaluation of the compiler in the logic. *) - -open preamble wordcountProgTheory compilationLib +open preamble wordcountProgTheory eval_cake_compile_ag32Lib val _ = new_theory"wordcountCompile"; -val wordcount_compiled = save_thm("wordcount_compiled", - compile_ag32 "wordcount" wordcount_prog_def); +Theorem wordcount_compiled = + eval_cake_compile_ag32 "" wordcount_prog_def "wordcount.S"; val _ = export_theory(); diff --git a/examples/compilation/x64/Holmakefile b/examples/compilation/x64/Holmakefile index 42715b6c9d..2e09fc1dfc 100644 --- a/examples/compilation/x64/Holmakefile +++ b/examples/compilation/x64/Holmakefile @@ -1,4 +1,4 @@ -INCLUDES = $(CAKEMLDIR)/misc $(CAKEMLDIR)/basis $(CAKEMLDIR)/compiler ../.. +INCLUDES = $(CAKEMLDIR)/misc $(CAKEMLDIR)/basis $(CAKEMLDIR)/compiler $(CAKEMLDIR)/cv_translator ../.. CLINE_OPTIONS = -j1 all: $(DEFAULT_TARGETS) README.md exec @@ -9,6 +9,10 @@ DIRS = $(wildcard */) README.md: $(CAKEMLDIR)/developers/readme_gen readmePrefix $(patsubst %,%readmePrefix,$(DIRS)) $(README_SOURCES) $(protect $(CAKEMLDIR)/developers/readme_gen) $(README_SOURCES) +ifdef POLY +HOLHEAP = $(CAKEMLDIR)/cv_translator/cake_compile_heap +endif + ifndef CC CC=gcc endif diff --git a/examples/compilation/x64/catCompileScript.sml b/examples/compilation/x64/catCompileScript.sml index d0688e9e54..659fa396a8 100644 --- a/examples/compilation/x64/catCompileScript.sml +++ b/examples/compilation/x64/catCompileScript.sml @@ -1,11 +1,11 @@ (* Compiles the cat example by evaluation inside the logic of HOL *) -open preamble compilationLib catProgTheory +open preamble catProgTheory eval_cake_compile_x64Lib val _ = new_theory "catCompile" -val cat_compiled = save_thm("cat_compiled", - compile_x64 "cat" cat_prog_def); +Theorem cat_compiled = + eval_cake_compile_x64 "" cat_prog_def "cat.S"; val _ = export_theory (); diff --git a/examples/compilation/x64/diffCompileScript.sml b/examples/compilation/x64/diffCompileScript.sml index 6a374e2e86..2805237476 100644 --- a/examples/compilation/x64/diffCompileScript.sml +++ b/examples/compilation/x64/diffCompileScript.sml @@ -1,11 +1,11 @@ (* Compiles the diff example by evaluation inside the logic of HOL *) -open preamble compilationLib diffProgTheory +open preamble diffProgTheory eval_cake_compile_x64Lib val _ = new_theory "diffCompile" -val diff_compiled = save_thm("diff_compiled", - compile_x64 "diff" diff_prog_def); +Theorem diff_compiled = + eval_cake_compile_x64 "" diff_prog_def "diff.S"; val _ = export_theory (); diff --git a/examples/compilation/x64/echoCompileScript.sml b/examples/compilation/x64/echoCompileScript.sml index 640fb0e4cb..5d38d8d00d 100644 --- a/examples/compilation/x64/echoCompileScript.sml +++ b/examples/compilation/x64/echoCompileScript.sml @@ -1,11 +1,11 @@ (* Compiles the echo example by evaluation inside the logic of HOL *) -open preamble compilationLib echoProgTheory +open preamble echoProgTheory eval_cake_compile_x64Lib val _ = new_theory "echoCompile" -val echo_compiled = save_thm("echo_compiled", - compile_x64 "echo" echo_prog_def); +Theorem echo_compiled = + eval_cake_compile_x64 "" echo_prog_def "echo.S"; val _ = export_theory (); diff --git a/examples/compilation/x64/grepCompileScript.sml b/examples/compilation/x64/grepCompileScript.sml index c37196b912..2a3d1e7786 100644 --- a/examples/compilation/x64/grepCompileScript.sml +++ b/examples/compilation/x64/grepCompileScript.sml @@ -1,11 +1,11 @@ (* Compiles the grep example by evaluation inside the logic of HOL *) -open preamble compilationLib grepProgTheory +open preamble grepProgTheory eval_cake_compile_x64Lib val _ = new_theory "grepCompile" -val grep_compiled = save_thm("grep_compiled", - compile_x64 "grep" grep_prog_def); +Theorem grep_compiled = + eval_cake_compile_x64 "" grep_prog_def "grep.S"; val _ = export_theory (); diff --git a/examples/compilation/x64/helloCompileScript.sml b/examples/compilation/x64/helloCompileScript.sml index 0fc2c19184..1108688994 100644 --- a/examples/compilation/x64/helloCompileScript.sml +++ b/examples/compilation/x64/helloCompileScript.sml @@ -1,13 +1,11 @@ (* Compiles the hello example by evaluation inside the logic of HOL *) -open preamble helloProgTheory compilationLib +open preamble helloProgTheory eval_cake_compile_x64Lib val _ = new_theory "helloCompile" -val _ = (output_ILs := SOME "hello"); - -val hello_compiled = save_thm("hello_compiled", - compile_x64 "hello" hello_prog_def); +Theorem hello_compiled = + eval_cake_compile_x64 "" hello_prog_def "hello.S"; val _ = export_theory (); diff --git a/examples/compilation/x64/helloErrCompileScript.sml b/examples/compilation/x64/helloErrCompileScript.sml index 069d00b8e4..3418fe6af7 100644 --- a/examples/compilation/x64/helloErrCompileScript.sml +++ b/examples/compilation/x64/helloErrCompileScript.sml @@ -1,11 +1,11 @@ (* Compiles the helloErr example by evaluation inside the logic of HOL *) -open preamble compilationLib helloErrProgTheory +open preamble helloErrProgTheory eval_cake_compile_x64Lib val _ = new_theory "helloErrCompile" -val helloErr_compiled = save_thm("helloErr_compiled", - compile_x64 "helloErr" helloErr_prog_def); +Theorem helloErr_compiled = + eval_cake_compile_x64 "" helloErr_prog_def "helloErr.S"; val _ = export_theory (); diff --git a/examples/compilation/x64/iocatCompileScript.sml b/examples/compilation/x64/iocatCompileScript.sml index 15b8da88d0..ec8e36adf7 100644 --- a/examples/compilation/x64/iocatCompileScript.sml +++ b/examples/compilation/x64/iocatCompileScript.sml @@ -1,11 +1,11 @@ (* Compiles the iocat example by evaluation inside the logic of HOL *) -open preamble compilationLib iocatProgTheory +open preamble iocatProgTheory eval_cake_compile_x64Lib val _ = new_theory "iocatCompile" -val cat_compiled = save_thm("iocat_compiled", - compile_x64 "iocat" cat_prog_def); +Theorem iocat_compiled = + eval_cake_compile_x64 "" cat_prog_def "iocat.S"; val _ = export_theory (); diff --git a/examples/compilation/x64/patchCompileScript.sml b/examples/compilation/x64/patchCompileScript.sml index de091520a2..b126ea2877 100644 --- a/examples/compilation/x64/patchCompileScript.sml +++ b/examples/compilation/x64/patchCompileScript.sml @@ -1,11 +1,11 @@ (* Compiles the patch example by evaluation inside the logic of HOL *) -open preamble compilationLib patchProgTheory +open preamble patchProgTheory eval_cake_compile_x64Lib val _ = new_theory "patchCompile" -val patch_compiled = save_thm("patch_compiled", - compile_x64 "patch" patch_prog_def); +Theorem patch_compiled = + eval_cake_compile_x64 "" patch_prog_def "patch.S"; val _ = export_theory (); diff --git a/examples/compilation/x64/proofs/catProofScript.sml b/examples/compilation/x64/proofs/catProofScript.sml index 816fba2c18..6fba963d02 100644 --- a/examples/compilation/x64/proofs/catProofScript.sml +++ b/examples/compilation/x64/proofs/catProofScript.sml @@ -17,7 +17,7 @@ val (cat_sem,cat_output) = cat_io_events_def |> SPEC_ALL |> UNDISCH_ALL |> CONJ_ val (cat_not_fail,cat_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail cat_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct cat_compiled + MATCH_MP compile_correct (cj 1 cat_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP cat_not_fail |> C MATCH_MP x64_backend_config_ok diff --git a/examples/compilation/x64/proofs/diffProofScript.sml b/examples/compilation/x64/proofs/diffProofScript.sml index 2122f6094c..b6696d3bba 100644 --- a/examples/compilation/x64/proofs/diffProofScript.sml +++ b/examples/compilation/x64/proofs/diffProofScript.sml @@ -17,7 +17,7 @@ val (diff_sem,diff_output) = diff_io_events_def |> SPEC_ALL |> UNDISCH |> CONJ_P val (diff_not_fail,diff_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail diff_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct diff_compiled + MATCH_MP compile_correct (cj 1 diff_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP diff_not_fail |> C MATCH_MP x64_backend_config_ok diff --git a/examples/compilation/x64/proofs/echoProofScript.sml b/examples/compilation/x64/proofs/echoProofScript.sml index 77f6f54f87..b1466cff35 100644 --- a/examples/compilation/x64/proofs/echoProofScript.sml +++ b/examples/compilation/x64/proofs/echoProofScript.sml @@ -18,7 +18,7 @@ val (echo_sem,echo_output) = echo_io_events_def |> SPEC_ALL |> UNDISCH |> CONJ_P val (echo_not_fail,echo_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail echo_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct echo_compiled + MATCH_MP compile_correct (cj 1 echo_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP echo_not_fail |> C MATCH_MP x64_backend_config_ok diff --git a/examples/compilation/x64/proofs/grepProofScript.sml b/examples/compilation/x64/proofs/grepProofScript.sml index 83cc581d2a..9d327a37d4 100644 --- a/examples/compilation/x64/proofs/grepProofScript.sml +++ b/examples/compilation/x64/proofs/grepProofScript.sml @@ -17,7 +17,7 @@ val (grep_sem,grep_output) = grep_io_events_def |> SPEC_ALL |> UNDISCH |> CONJ_P val (grep_not_fail,grep_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail grep_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct grep_compiled + MATCH_MP compile_correct (cj 1 grep_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP grep_not_fail |> C MATCH_MP x64_backend_config_ok diff --git a/examples/compilation/x64/proofs/helloErrProofScript.sml b/examples/compilation/x64/proofs/helloErrProofScript.sml index 84e5fd7645..e3e0c1f012 100644 --- a/examples/compilation/x64/proofs/helloErrProofScript.sml +++ b/examples/compilation/x64/proofs/helloErrProofScript.sml @@ -19,7 +19,7 @@ val (helloErr_not_fail,helloErr_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail helloErr_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct helloErr_compiled + MATCH_MP compile_correct (cj 1 helloErr_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP helloErr_not_fail |> C MATCH_MP x64_backend_config_ok diff --git a/examples/compilation/x64/proofs/helloProofScript.sml b/examples/compilation/x64/proofs/helloProofScript.sml index 10798fe79b..df7cfe24cf 100644 --- a/examples/compilation/x64/proofs/helloProofScript.sml +++ b/examples/compilation/x64/proofs/helloProofScript.sml @@ -18,7 +18,7 @@ val (hello_sem,hello_output) = hello_io_events_def |> SPEC_ALL |> UNDISCH |> CON val (hello_not_fail,hello_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail hello_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct hello_compiled + MATCH_MP compile_correct (cj 1 hello_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP hello_not_fail |> C MATCH_MP x64_backend_config_ok diff --git a/examples/compilation/x64/proofs/iocatProofScript.sml b/examples/compilation/x64/proofs/iocatProofScript.sml index 29d526cd57..750aa2f961 100644 --- a/examples/compilation/x64/proofs/iocatProofScript.sml +++ b/examples/compilation/x64/proofs/iocatProofScript.sml @@ -18,7 +18,7 @@ val (cat_sem,cat_output) = cat_io_events_def |> SPEC_ALL |> UNDISCH_ALL |> CONJ_ val (cat_not_fail,cat_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail cat_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct iocat_compiled + MATCH_MP compile_correct (cj 1 iocat_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP cat_not_fail |> C MATCH_MP x64_backend_config_ok diff --git a/examples/compilation/x64/proofs/patchProofScript.sml b/examples/compilation/x64/proofs/patchProofScript.sml index 2a3adc0eb0..4c847ac9b0 100644 --- a/examples/compilation/x64/proofs/patchProofScript.sml +++ b/examples/compilation/x64/proofs/patchProofScript.sml @@ -17,7 +17,7 @@ val (patch_sem,patch_output) = patch_io_events_def |> SPEC_ALL |> UNDISCH |> CON val (patch_not_fail,patch_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail patch_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct patch_compiled + MATCH_MP compile_correct (cj 1 patch_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP patch_not_fail |> C MATCH_MP x64_backend_config_ok diff --git a/examples/compilation/x64/proofs/sortProofScript.sml b/examples/compilation/x64/proofs/sortProofScript.sml index f30445c376..83dedb3821 100644 --- a/examples/compilation/x64/proofs/sortProofScript.sml +++ b/examples/compilation/x64/proofs/sortProofScript.sml @@ -17,7 +17,7 @@ val (sort_sem,sort_output) = sort_io_events_def |> SPEC_ALL |> UNDISCH |> CONJ_P val (sort_not_fail,sort_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail sort_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct sort_compiled + MATCH_MP compile_correct (cj 1 sort_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP sort_not_fail |> C MATCH_MP x64_backend_config_ok diff --git a/examples/compilation/x64/sortCompileScript.sml b/examples/compilation/x64/sortCompileScript.sml index 8e96097cc1..2958fb118f 100644 --- a/examples/compilation/x64/sortCompileScript.sml +++ b/examples/compilation/x64/sortCompileScript.sml @@ -1,11 +1,11 @@ (* Compiles the sort example by evaluation inside the logic of HOL *) -open preamble compilationLib sortProgTheory +open preamble sortProgTheory eval_cake_compile_x64Lib val _ = new_theory "sortCompile" -val sort_compiled = save_thm("sort_compiled", - compile_x64 "sort" sort_prog_def); +Theorem sort_compiled = + eval_cake_compile_x64 "" sort_prog_def "sort.S"; val _ = export_theory (); diff --git a/examples/deflate/translation/compilation/Holmakefile b/examples/deflate/translation/compilation/Holmakefile index 81783522ff..1ac6e30314 100644 --- a/examples/deflate/translation/compilation/Holmakefile +++ b/examples/deflate/translation/compilation/Holmakefile @@ -1,4 +1,4 @@ -INCLUDES = $(CAKEMLDIR)/examples/deflate $(CAKEMLDIR)/examples/deflate/translation $(CAKEMLDIR)/compiler +INCLUDES = $(CAKEMLDIR)/examples/deflate $(CAKEMLDIR)/examples/deflate/translation $(CAKEMLDIR)/compiler $(CAKEMLDIR)/cv_translator all: $(DEFAULT_TARGETS) README.md @@ -9,6 +9,10 @@ DIRS = $(wildcard */) README.md: $(CAKEMLDIR)/developers/readme_gen readmePrefix $(README_SOURCES) $(protect $(CAKEMLDIR)/developers/readme_gen) $(README_SOURCES) +ifdef POLY +HOLHEAP = $(CAKEMLDIR)/cv_translator/cake_compile_heap +endif + compression.S: compressionCompileScript.sml decompression.S: decompressionCompileScript.sml deflateEncode.S: deflateEncodeCompileScript.sml diff --git a/examples/deflate/translation/compilation/compressionCompileScript.sml b/examples/deflate/translation/compilation/compressionCompileScript.sml index 3744f0946b..51c7b665d0 100644 --- a/examples/deflate/translation/compilation/compressionCompileScript.sml +++ b/examples/deflate/translation/compilation/compressionCompileScript.sml @@ -1,12 +1,11 @@ (* Compiles the simple compression schema *) - -open preamble compilationLib compressionProgTheory; +open preamble compressionProgTheory eval_cake_compile_x64Lib; val _ = new_theory "compressionCompile" Theorem compression_compiled = - compile_x64 "compression" compression_prog_def; + eval_cake_compile_x64 "" compression_prog_def "compression.S"; val _ = export_theory (); diff --git a/examples/deflate/translation/compilation/decompressionCompileScript.sml b/examples/deflate/translation/compilation/decompressionCompileScript.sml index ef2b6797ec..481c92c673 100644 --- a/examples/deflate/translation/compilation/decompressionCompileScript.sml +++ b/examples/deflate/translation/compilation/decompressionCompileScript.sml @@ -1,12 +1,11 @@ (* Compiles the simple decompression schema *) - -open preamble compilationLib decompressionProgTheory; +open preamble decompressionProgTheory eval_cake_compile_x64Lib; val _ = new_theory "decompressionCompile" Theorem decompression_compiled = - compile_x64 "decompression" decompression_prog_def; + eval_cake_compile_x64 "" decompression_prog_def "decompression.S"; val _ = export_theory (); diff --git a/examples/deflate/translation/compilation/deflateDecodeCompileScript.sml b/examples/deflate/translation/compilation/deflateDecodeCompileScript.sml index 46659232a9..504940d07c 100644 --- a/examples/deflate/translation/compilation/deflateDecodeCompileScript.sml +++ b/examples/deflate/translation/compilation/deflateDecodeCompileScript.sml @@ -1,12 +1,11 @@ (* Compiles the binary for the Deflate decoder *) - -open preamble compilationLib deflateDecodeProgTheory; +open preamble deflateDecodeProgTheory eval_cake_compile_x64Lib; val _ = new_theory "deflateDecodeCompile" Theorem deflateDecode_compiled = - compile_x64 "deflateDecode" deflateDecode_prog_def; + eval_cake_compile_x64 "" deflateDecode_prog_def "deflateDecode.S"; val _ = export_theory (); diff --git a/examples/deflate/translation/compilation/deflateEncodeCompileScript.sml b/examples/deflate/translation/compilation/deflateEncodeCompileScript.sml index 50ac0c0b32..7d35363370 100644 --- a/examples/deflate/translation/compilation/deflateEncodeCompileScript.sml +++ b/examples/deflate/translation/compilation/deflateEncodeCompileScript.sml @@ -1,12 +1,11 @@ (* Compiles the binary for the Deflate Encoder *) - -open preamble compilationLib deflateEncodeProgTheory; +open preamble deflateEncodeProgTheory eval_cake_compile_x64Lib; val _ = new_theory "deflateEncodeCompile" Theorem deflateEncode_compiled = - compile_x64 "deflateEncode" deflateEncode_prog_def; + eval_cake_compile_x64 "" deflateEncode_prog_def "deflateEncode.S"; val _ = export_theory (); diff --git a/examples/lpr_checker/array/compilation/Holmakefile b/examples/lpr_checker/array/compilation/Holmakefile index 8e7cea16c9..bef537dbe0 100644 --- a/examples/lpr_checker/array/compilation/Holmakefile +++ b/examples/lpr_checker/array/compilation/Holmakefile @@ -1,4 +1,4 @@ -INCLUDES = $(CAKEMLDIR)/misc $(CAKEMLDIR)/basis $(CAKEMLDIR)/compiler .. +INCLUDES = $(CAKEMLDIR)/misc $(CAKEMLDIR)/basis $(CAKEMLDIR)/compiler $(CAKEMLDIR)/cv_translator .. CLINE_OPTIONS = all: $(DEFAULT_TARGETS) README.md @@ -9,6 +9,10 @@ DIRS = $(wildcard */) README.md: $(CAKEMLDIR)/developers/readme_gen readmePrefix $(patsubst %,%readmePrefix,$(DIRS)) $(README_SOURCES) $(protect $(CAKEMLDIR)/developers/readme_gen) $(README_SOURCES) +ifdef POLY +HOLHEAP = $(CAKEMLDIR)/cv_translator/cake_compile_heap +endif + cake_lpr.S: *lpr_arrayCompileScript.sml cake_lpr_arm8.S: *lpr_arrayCompileARM8Script.sml cake_lpr_ramsey.S: *lpr_arrayRamseyCompileScript.sml diff --git a/examples/lpr_checker/array/compilation/lpr_arrayCompileARM8Script.sml b/examples/lpr_checker/array/compilation/lpr_arrayCompileARM8Script.sml index 7fef52b730..5c653bd294 100644 --- a/examples/lpr_checker/array/compilation/lpr_arrayCompileARM8Script.sml +++ b/examples/lpr_checker/array/compilation/lpr_arrayCompileARM8Script.sml @@ -1,11 +1,11 @@ (* Compiles the lpr example by evaluation inside the logic of HOL *) -open preamble compilationLib lpr_arrayFullProgTheory +open preamble lpr_arrayFullProgTheory eval_cake_compile_arm8Lib val _ = new_theory "lpr_arrayCompileARM8" -val lpr_array_compiled = save_thm("lpr_array_compiled", - compile_arm8 "cake_lpr_arm8" check_unsat_prog_def); +Theorem lpr_array_compiled = + eval_cake_compile_arm8 "" check_unsat_prog_def "cake_lpr_arm8.S"; val _ = export_theory (); diff --git a/examples/lpr_checker/array/compilation/lpr_arrayCompileScript.sml b/examples/lpr_checker/array/compilation/lpr_arrayCompileScript.sml index e8918b2031..f430500ba0 100644 --- a/examples/lpr_checker/array/compilation/lpr_arrayCompileScript.sml +++ b/examples/lpr_checker/array/compilation/lpr_arrayCompileScript.sml @@ -1,11 +1,11 @@ (* Compiles the lpr example by evaluation inside the logic of HOL *) -open preamble compilationLib lpr_arrayFullProgTheory +open preamble lpr_arrayFullProgTheory eval_cake_compile_x64Lib val _ = new_theory "lpr_arrayCompile" -val lpr_array_compiled = save_thm("lpr_array_compiled", - compile_x64 "cake_lpr" check_unsat_prog_def); +Theorem lpr_array_compiled = + eval_cake_compile_x64 "" check_unsat_prog_def "cake_lpr.S"; val _ = export_theory (); diff --git a/examples/lpr_checker/array/compilation/lpr_arrayPackingCompileScript.sml b/examples/lpr_checker/array/compilation/lpr_arrayPackingCompileScript.sml index 666a06a065..b771fc5b1c 100644 --- a/examples/lpr_checker/array/compilation/lpr_arrayPackingCompileScript.sml +++ b/examples/lpr_checker/array/compilation/lpr_arrayPackingCompileScript.sml @@ -1,11 +1,11 @@ (* Compiles the packing example by evaluation inside the logic of HOL *) -open preamble compilationLib lpr_arrayPackingProgTheory +open preamble lpr_arrayPackingProgTheory eval_cake_compile_x64Lib val _ = new_theory "lpr_arrayPackingCompile" -val lpr_array_compiled = save_thm("lpr_packing_compiled", - compile_x64 "cake_direct" main_prog_def); +Theorem lpr_packing_compiled = + eval_cake_compile_x64 "" main_prog_def "cake_direct.S"; val _ = export_theory (); diff --git a/examples/lpr_checker/array/compilation/lpr_arrayRamseyCompileScript.sml b/examples/lpr_checker/array/compilation/lpr_arrayRamseyCompileScript.sml index ebcd4b0e4c..2555d9d3a2 100644 --- a/examples/lpr_checker/array/compilation/lpr_arrayRamseyCompileScript.sml +++ b/examples/lpr_checker/array/compilation/lpr_arrayRamseyCompileScript.sml @@ -1,11 +1,11 @@ (* Compiles the lpr example by evaluation inside the logic of HOL *) -open preamble compilationLib lpr_arrayRamseyProgTheory +open preamble lpr_arrayRamseyProgTheory eval_cake_compile_x64Lib val _ = new_theory "lpr_arrayRamseyCompile" -val lpr_array_compiled = save_thm("lpr_ramsey_compiled", - compile_x64 "cake_lpr_ramsey" check_unsat_prog_def); +Theorem lpr_ramsey_compiled = + eval_cake_compile_x64 "" check_unsat_prog_def "cake_lpr_ramsey.S"; val _ = export_theory (); diff --git a/examples/lpr_checker/array/compilation/proofs/lpr_arrayPackingProofScript.sml b/examples/lpr_checker/array/compilation/proofs/lpr_arrayPackingProofScript.sml index 731036804c..bafa77cbc9 100644 --- a/examples/lpr_checker/array/compilation/proofs/lpr_arrayPackingProofScript.sml +++ b/examples/lpr_checker/array/compilation/proofs/lpr_arrayPackingProofScript.sml @@ -20,7 +20,7 @@ val (main_sem,main_output) = main_io_events_def |> SPEC_ALL |> UNDISCH |> SIMP_R val (main_not_fail,main_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail main_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct lpr_packing_compiled + MATCH_MP compile_correct (cj 1 lpr_packing_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP main_not_fail |> C MATCH_MP x64_backend_config_ok @@ -53,7 +53,7 @@ val installed_x64_def = Define ` `; val main_code_def = Define ` - main_code = (code, data, config) + main_code = (code, data, info) `; (* A standard run of packing satisfying all the default assumptions *) diff --git a/examples/lpr_checker/array/compilation/proofs/lpr_arrayProofScript.sml b/examples/lpr_checker/array/compilation/proofs/lpr_arrayProofScript.sml index 16c21896c3..bcab24a275 100644 --- a/examples/lpr_checker/array/compilation/proofs/lpr_arrayProofScript.sml +++ b/examples/lpr_checker/array/compilation/proofs/lpr_arrayProofScript.sml @@ -19,7 +19,7 @@ val (check_unsat_sem,check_unsat_output) = check_unsat_io_events_def |> SPEC_ALL val (check_unsat_not_fail,check_unsat_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail check_unsat_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct lpr_array_compiled + MATCH_MP compile_correct (cj 1 lpr_array_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP check_unsat_not_fail |> C MATCH_MP x64_backend_config_ok @@ -52,7 +52,7 @@ val installed_x64_def = Define ` `; val check_unsat_code_def = Define ` - check_unsat_code = (code, data, config) + check_unsat_code = (code, data, info) `; (* A standard run of cake_lpr satisfying all the default assumptions *) diff --git a/examples/lpr_checker/array/compilation/proofs/lpr_arrayRamseyProofScript.sml b/examples/lpr_checker/array/compilation/proofs/lpr_arrayRamseyProofScript.sml index 9322c8b1a3..6330a1f348 100644 --- a/examples/lpr_checker/array/compilation/proofs/lpr_arrayRamseyProofScript.sml +++ b/examples/lpr_checker/array/compilation/proofs/lpr_arrayRamseyProofScript.sml @@ -20,7 +20,7 @@ val (check_unsat_sem,check_unsat_output) = check_unsat_io_events_def |> SPEC_ALL val (check_unsat_not_fail,check_unsat_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail check_unsat_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct lpr_ramsey_compiled + MATCH_MP compile_correct (cj 1 lpr_ramsey_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP check_unsat_not_fail |> C MATCH_MP x64_backend_config_ok @@ -53,7 +53,7 @@ val installed_x64_def = Define ` `; val check_unsat_code_def = Define ` - check_unsat_code = (code, data, config) + check_unsat_code = (code, data, info) `; (* A standard run of ramsey satisfying all the default assumptions *) diff --git a/examples/lpr_checker/array/compilation/proofsARM8/Holmakefile b/examples/lpr_checker/array/compilation/proofsARM8/Holmakefile index 4ad9d618ea..ea2c5bc033 100644 --- a/examples/lpr_checker/array/compilation/proofsARM8/Holmakefile +++ b/examples/lpr_checker/array/compilation/proofsARM8/Holmakefile @@ -13,3 +13,9 @@ README_SOURCES = $(wildcard *Script.sml) $(wildcard *Lib.sml) $(wildcard *Syntax DIRS = $(wildcard */) README.md: $(CAKEMLDIR)/developers/readme_gen readmePrefix $(patsubst %,%readmePrefix,$(DIRS)) $(README_SOURCES) $(protect $(CAKEMLDIR)/developers/readme_gen) $(README_SOURCES) + +ifdef POLY + +HOLHEAP=$(HOLDIR)/examples/l3-machine-code/arm8/asl-equiv/l3-equivalence-heap + +endif diff --git a/examples/lpr_checker/array/compilation/proofsARM8/lpr_arrayProofARM8Script.sml b/examples/lpr_checker/array/compilation/proofsARM8/lpr_arrayProofARM8Script.sml index 6cf17fb7a1..f4d274bf5a 100644 --- a/examples/lpr_checker/array/compilation/proofsARM8/lpr_arrayProofARM8Script.sml +++ b/examples/lpr_checker/array/compilation/proofsARM8/lpr_arrayProofARM8Script.sml @@ -20,7 +20,7 @@ val (check_unsat_sem,check_unsat_output) = check_unsat_io_events_def |> SPEC_ALL val (check_unsat_not_fail,check_unsat_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail check_unsat_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct lpr_array_compiled + MATCH_MP compile_correct (cj 1 lpr_array_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP check_unsat_not_fail |> C MATCH_MP arm8_configProofTheory.arm8_backend_config_ok @@ -53,7 +53,7 @@ val installed_arm8_asl_def = Define ` `; val check_unsat_code_def = Define ` - check_unsat_code = (code, data, config) + check_unsat_code = (code, data, info) `; (* A standard run of cake_lpr satisfying all the default assumptions *) diff --git a/examples/opentheory/compilation/Holmakefile b/examples/opentheory/compilation/Holmakefile index 7099bc9c3e..e32bff28dd 100644 --- a/examples/opentheory/compilation/Holmakefile +++ b/examples/opentheory/compilation/Holmakefile @@ -1,4 +1,4 @@ -INCLUDES = $(CAKEMLDIR)/misc $(CAKEMLDIR)/compiler $(CAKEMLDIR)/examples/cost .. +INCLUDES = $(CAKEMLDIR)/misc $(CAKEMLDIR)/cv_translator $(CAKEMLDIR)/compiler $(CAKEMLDIR)/examples/cost .. all: $(DEFAULT_TARGETS) README.md .PHONY: all @@ -8,6 +8,10 @@ DIRS = $(wildcard */) README.md: $(CAKEMLDIR)/developers/readme_gen readmePrefix $(patsubst %,%readmePrefix,$(DIRS)) $(README_SOURCES) $(protect $(CAKEMLDIR)/developers/readme_gen) $(README_SOURCES) +ifdef POLY +HOLHEAP = $(CAKEMLDIR)/cv_translator/cake_compile_heap +endif + reader: reader.S $(CAKEMLDIR)/basis/basis_ffi.o $(CC) $< $(protect $(CAKEMLDIR)/basis/basis_ffi.o) $(GCCFLAGS) -o $@ diff --git a/examples/opentheory/compilation/ag32/Holmakefile b/examples/opentheory/compilation/ag32/Holmakefile index 3410fc54dc..da39b43dc2 100644 --- a/examples/opentheory/compilation/ag32/Holmakefile +++ b/examples/opentheory/compilation/ag32/Holmakefile @@ -1,4 +1,4 @@ -INCLUDES = $(CAKEMLDIR)/misc $(CAKEMLDIR)/basis $(CAKEMLDIR)/compiler\ +INCLUDES = $(CAKEMLDIR)/misc $(CAKEMLDIR)/basis $(CAKEMLDIR)/cv_translator $(CAKEMLDIR)/compiler\ ../.. all: $(DEFAULT_TARGETS) README.md @@ -8,3 +8,7 @@ README_SOURCES = $(wildcard *Script.sml) $(wildcard *Lib.sml) $(wildcard *Syntax DIRS = $(wildcard */) README.md: $(CAKEMLDIR)/developers/readme_gen readmePrefix $(patsubst %,%readmePrefix,$(DIRS)) $(README_SOURCES) $(protect $(CAKEMLDIR)/developers/readme_gen) $(README_SOURCES) + +ifdef POLY +HOLHEAP = $(CAKEMLDIR)/cv_translator/cake_compile_heap +endif diff --git a/examples/opentheory/compilation/ag32/proofs/readerProgProofScript.sml b/examples/opentheory/compilation/ag32/proofs/readerProgProofScript.sml index 317b0d3474..b9876a5e6a 100644 --- a/examples/opentheory/compilation/ag32/proofs/readerProgProofScript.sml +++ b/examples/opentheory/compilation/ag32/proofs/readerProgProofScript.sml @@ -21,30 +21,34 @@ val reader_io_events_def = val (reader_sem,reader_output) = reader_io_events_def |> SPEC_ALL |> UNDISCH |> CONJ_PAIR val (reader_not_fail,reader_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail reader_sem |> CONJ_PAIR -val ffi_names = - ``config.lab_conf.ffi_names`` - |> (REWRITE_CONV[readerCompileTheory.config_def] THENC EVAL) +val ffinames_to_string_list_def = backendTheory.ffinames_to_string_list_def; -val LENGTH_code = - ``LENGTH code`` - |> (REWRITE_CONV[readerCompileTheory.code_def] THENC listLib.LENGTH_CONV) +Theorem extcalls_ffi_names: + extcalls info.lab_conf.ffi_names = ffis +Proof + rewrite_tac [reader_compiled] + \\ qspec_tac (‘info.lab_conf.ffi_names’,‘xs’) \\ Cases + \\ gvs [extcalls_def,ffinames_to_string_list_def,libTheory.the_def] + \\ Induct_on ‘x’ + \\ gvs [extcalls_def,ffinames_to_string_list_def,libTheory.the_def] + \\ Cases \\ gvs [extcalls_def,ffinames_to_string_list_def,libTheory.the_def] +QED -val LENGTH_data = - ``LENGTH data`` - |> (REWRITE_CONV[readerCompileTheory.data_def] THENC listLib.LENGTH_CONV) +val ffis = ffis_def |> CONV_RULE (RAND_CONV EVAL); +val ffi_names = extcalls_ffi_names |> SRULE [ffis] -val shmem = - ``config.lab_conf.shmem_extra`` - |> (REWRITE_CONV[readerCompileTheory.config_def] THENC EVAL) +val LENGTH_code = “LENGTH code” |> SCONV [reader_compiled]; +val LENGTH_data = “LENGTH data” |> SCONV [reader_compiled]; +val shmem = “info.lab_conf.shmem_extra” |> SCONV [reader_compiled]; Overload reader_machine_config = - ``ag32_machine_config (extcalls config.lab_conf.ffi_names) (LENGTH code) (LENGTH data)`` + “ag32_machine_config (extcalls info.lab_conf.ffi_names) (LENGTH code) (LENGTH data)” Theorem target_state_rel_reader_start_asm_state: SUM (MAP strlen cl) + LENGTH cl ≤ cline_size ∧ LENGTH inp ≤ stdin_size ∧ - is_ag32_init_state (init_memory code data (extcalls config.lab_conf.ffi_names) (cl,inp)) ms ⇒ - ∃n. target_state_rel ag32_target (init_asm_state code data (extcalls config.lab_conf.ffi_names) (cl,inp)) (FUNPOW Next n ms) ∧ + is_ag32_init_state (init_memory code data (extcalls info.lab_conf.ffi_names) (cl,inp)) ms ⇒ + ∃n. target_state_rel ag32_target (init_asm_state code data (extcalls info.lab_conf.ffi_names) (cl,inp)) (FUNPOW Next n ms) ∧ ((FUNPOW Next n ms).io_events = ms.io_events) ∧ (∀x. x ∉ (ag32_startup_addresses) ⇒ ((FUNPOW Next n ms).MEM x = ms.MEM x)) @@ -53,7 +57,7 @@ Proof \\ drule (GEN_ALL init_asm_state_RTC_asm_step) \\ disch_then drule \\ simp_tac std_ss [] - \\ disch_then(qspecl_then[`code`,`data`,`extcalls config.lab_conf.ffi_names`]mp_tac) + \\ disch_then(qspecl_then[`code`,`data`,`extcalls info.lab_conf.ffi_names`]mp_tac) \\ impl_tac >- ( EVAL_TAC>> fs[ffi_names,LENGTH_data,LENGTH_code,extcalls_def]) \\ strip_tac \\ drule (GEN_ALL target_state_rel_ag32_init) @@ -70,7 +74,7 @@ val reader_startup_clock_def = |> SIMP_RULE bool_ss [GSYM RIGHT_EXISTS_IMP_THM,SKOLEM_THM]); val compile_correct_applied = - MATCH_MP compile_correct reader_compiled + MATCH_MP compile_correct (cj 1 reader_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP reader_not_fail |> C MATCH_MP ag32_backend_config_ok @@ -82,29 +86,33 @@ val compile_correct_applied = |> Q.GEN`cbspace` |> Q.SPEC`0` |> Q.GEN`data_sp` |> Q.SPEC`0` -Triviality to_MAP_ExtCall: - [ExtCall n] = MAP ExtCall [n] ∧ - (ExtCall n::MAP ExtCall ns) = MAP ExtCall (n::ns) -Proof - fs [] -QED - Theorem reader_installed: SUM (MAP strlen cl) + LENGTH cl ≤ cline_size ∧ LENGTH inp ≤ stdin_size ∧ - is_ag32_init_state (init_memory code data (extcalls config.lab_conf.ffi_names) (cl,inp)) ms0 ⇒ - installed code 0 data 0 config.lab_conf.ffi_names + is_ag32_init_state (init_memory code data (extcalls info.lab_conf.ffi_names) (cl,inp)) ms0 ⇒ + installed code 0 data 0 info.lab_conf.ffi_names (heap_regs ag32_backend_config.stack_conf.reg_names) - (reader_machine_config) config.lab_conf.shmem_extra + (reader_machine_config) info.lab_conf.shmem_extra (FUNPOW Next (reader_startup_clock ms0 inp cl) ms0) Proof rewrite_tac[ffi_names, extcalls_def, shmem] \\ strip_tac - \\ rewrite_tac [to_MAP_ExtCall] + \\ qmatch_asmsub_abbrev_tac ‘init_memory _ _ ff’ + \\ qmatch_goalsub_abbrev_tac ‘installed _ _ _ _ dd’ + \\ ‘dd = SOME (MAP ExtCall ff)’ by + (unabbrev_all_tac + \\ assume_tac (cj 1 reader_compiled) + \\ drule ag32_configProofTheory.compile_imp_ffi_names + \\ gvs [reader_compiled] + \\ gvs [GSYM reader_compiled,ffis] + \\ simp [backendTheory.set_oracle_def, + ag32_configTheory.ag32_backend_config_def]) + \\ asm_rewrite_tac [] \\ irule ag32_installed \\ drule reader_startup_clock_def \\ disch_then drule \\ rewrite_tac[ffi_names, extcalls_def] + \\ unabbrev_all_tac \\ disch_then drule \\ strip_tac \\ simp[] @@ -173,7 +181,7 @@ Proof \\ simp [RIGHT_EXISTS_AND_THM] \\ simp [readerProofTheory.reader_main_def, readerProofTheory.read_stdin_def] - \\ qpat_x_assum ‘_ = init_reader _ _’ (assume_tac o SYM) +(* \\ qpat_x_assum ‘_ = init_reader _ _’ (assume_tac o SYM) *) \\ simp [all_lines_stdin_fs] \\ (conj_tac >- simp [fsFFIPropsTheory.inFS_fname_def, stdin_fs_def]) \\ (conj_tac >- simp [stdin_fs_def]) @@ -209,7 +217,7 @@ Proof \\ simp [RIGHT_EXISTS_AND_THM] \\ simp [readerProofTheory.reader_main_def, readerProofTheory.read_stdin_def] - \\ qpat_x_assum ‘_ = init_reader _ _’ (assume_tac o SYM) +(*\\ qpat_x_assum ‘_ = init_reader _ _’ (assume_tac o SYM) *) \\ simp [all_lines_stdin_fs] \\ (conj_tac >- simp [fsFFIPropsTheory.inFS_fname_def, stdin_fs_def]) \\ (conj_tac >- simp [stdin_fs_def]) @@ -232,7 +240,7 @@ Theorem reader_ag32_next: LENGTH inp <= stdin_size /\ wfcl cl /\ (LENGTH cl = 1) /\ - is_ag32_init_state (init_memory code data (extcalls config.lab_conf.ffi_names) (cl,inp)) ms0 + is_ag32_init_state (init_memory code data (extcalls info.lab_conf.ffi_names) (cl,inp)) ms0 ==> ?k1. !k. k1 <= k ==> let ms = FUNPOW Next k ms0 in diff --git a/examples/opentheory/compilation/ag32/readerCompileScript.sml b/examples/opentheory/compilation/ag32/readerCompileScript.sml index d40ef433a8..4e3f6e8867 100644 --- a/examples/opentheory/compilation/ag32/readerCompileScript.sml +++ b/examples/opentheory/compilation/ag32/readerCompileScript.sml @@ -2,11 +2,11 @@ In-logic compilation of the OPenTHeory article checker to the Silver ISA. *) -open preamble compilationLib readerProgTheory +open preamble readerProgTheory eval_cake_compile_ag32Lib val _ = new_theory "readerCompile" -val reader_compiled = save_thm ("reader_compiled", - compile_ag32 "reader" reader_prog_def); +Theorem reader_compiled = + eval_cake_compile_ag32 "" reader_prog_def "reader.S"; val _ = export_theory (); diff --git a/examples/opentheory/compilation/proofs/readerProgProofScript.sml b/examples/opentheory/compilation/proofs/readerProgProofScript.sml index 479905d2a8..ec76e0a484 100644 --- a/examples/opentheory/compilation/proofs/readerProgProofScript.sml +++ b/examples/opentheory/compilation/proofs/readerProgProofScript.sml @@ -21,7 +21,7 @@ val (reader_not_fail, reader_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail reader_sem |> CONJ_PAIR; val compile_correct_applied = - MATCH_MP compile_correct reader_compiled + MATCH_MP compile_correct (cj 1 reader_compiled) |> SIMP_RULE (srw_ss()) [LET_THM, ml_progTheory.init_state_env_thm, GSYM AND_IMP_INTRO] |> C MATCH_MP reader_not_fail @@ -57,7 +57,7 @@ Definition installed_x64_def: End Definition reader_code_def: - reader_code = (code, data, config) + reader_code = (code, data, info) End val _ = Parse.hide "mem"; @@ -91,4 +91,3 @@ Proof QED val _ = export_theory (); - diff --git a/examples/opentheory/compilation/readerCompileScript.sml b/examples/opentheory/compilation/readerCompileScript.sml index 144ae2a831..73006c2268 100644 --- a/examples/opentheory/compilation/readerCompileScript.sml +++ b/examples/opentheory/compilation/readerCompileScript.sml @@ -1,16 +1,16 @@ (* Compiles the OpenTheory article checker in the logic. *) -open preamble compilationLib readerProgTheory +open preamble readerProgTheory eval_cake_compile_x64Lib open x64_configTheory -open costLib +(* open costLib *) val _ = new_theory "readerCompile" -val reader_compiled = save_thm("reader_compiled", - compile_x64 "reader" reader_prog_def); +Theorem reader_compiled = + eval_cake_compile_x64 "" reader_prog_def "reader.S"; -(* the following stores a pretty printer the dataLang program in a textfile *) +(* the following stores a pretty printer the dataLang program in a textfile Overload monad_unitbind[local] = ``data_monad$bind`` Overload return[local] = ``data_monad$return`` @@ -19,4 +19,6 @@ val _ = monadsyntax.temp_add_monadsyntax() val _ = install_naming_overloads "readerCompile"; val _ = write_to_file (fetch "-" "data_prog_def"); +*) + val _ = export_theory (); diff --git a/examples/pseudo_bool/array/compilation/Holmakefile b/examples/pseudo_bool/array/compilation/Holmakefile index 3d54715276..94260fb901 100644 --- a/examples/pseudo_bool/array/compilation/Holmakefile +++ b/examples/pseudo_bool/array/compilation/Holmakefile @@ -1,4 +1,4 @@ -INCLUDES = $(CAKEMLDIR)/misc $(CAKEMLDIR)/basis $(CAKEMLDIR)/compiler .. $(CAKEMLDIR)/unverified/sexpr-bootstrap $(CAKEMLDIR)/compiler/parsing +INCLUDES = $(CAKEMLDIR)/misc $(CAKEMLDIR)/basis $(CAKEMLDIR)/compiler $(CAKEMLDIR)/cv_translator .. $(CAKEMLDIR)/unverified/sexpr-bootstrap $(CAKEMLDIR)/compiler/parsing CLINE_OPTIONS = -j4 all: $(DEFAULT_TARGETS) README.md @@ -9,6 +9,10 @@ DIRS = $(wildcard */) README.md: $(CAKEMLDIR)/developers/readme_gen readmePrefix $(patsubst %,%readmePrefix,$(DIRS)) $(README_SOURCES) $(protect $(CAKEMLDIR)/developers/readme_gen) $(README_SOURCES) +ifdef POLY +HOLHEAP = $(CAKEMLDIR)/cv_translator/cake_compile_heap +endif + ifndef CC CC=gcc endif diff --git a/examples/pseudo_bool/array/compilation/README.md b/examples/pseudo_bool/array/compilation/README.md index 0578ae8601..6970127541 100644 --- a/examples/pseudo_bool/array/compilation/README.md +++ b/examples/pseudo_bool/array/compilation/README.md @@ -18,6 +18,9 @@ Compiles the PB checker example by evaluation inside the logic of HOL [proofs](proofs): Prove end-to-end correctness theorem for PB checker with arrays +[proofsARM8](proofsARM8): +Prove end-to-end correctness theorem for PB checker with arrays + [subgraph_isoCompileScript.sml](subgraph_isoCompileScript.sml): Compiles the encoder diff --git a/examples/pseudo_bool/array/compilation/cliqueCompileScript.sml b/examples/pseudo_bool/array/compilation/cliqueCompileScript.sml index 82d887a838..fafa0cd3eb 100644 --- a/examples/pseudo_bool/array/compilation/cliqueCompileScript.sml +++ b/examples/pseudo_bool/array/compilation/cliqueCompileScript.sml @@ -1,19 +1,15 @@ (* Compiles the max clique + PB checker *) -open preamble compilationLib cliqueProgTheory +open preamble cliqueProgTheory eval_cake_compile_x64Lib + eval_cake_compile_arm8Lib val _ = new_theory "cliqueCompile" -val _ = intermediate_prog_prefix := "arm8_" +Theorem clique_compiled = + eval_cake_compile_x64 "" main_prog_def "cake_pb_clique.S"; -val clique_compiled_arm8 = save_thm("clique_compiled_arm8", - compile_arm8 "cake_pb_clique_arm8" main_prog_def); - -(* Default has no prefix *) -val _ = intermediate_prog_prefix := "" - -val clique_compiled = save_thm("clique_compiled", - compile_x64 "cake_pb_clique" main_prog_def); +Theorem clique_compiled_arm8 = + eval_cake_compile_arm8 "arm8_" main_prog_def "cake_pb_clique_arm8.S"; val _ = export_theory (); diff --git a/examples/pseudo_bool/array/compilation/cnfCompileScript.sml b/examples/pseudo_bool/array/compilation/cnfCompileScript.sml index 4adb965071..c0ac613de4 100644 --- a/examples/pseudo_bool/array/compilation/cnfCompileScript.sml +++ b/examples/pseudo_bool/array/compilation/cnfCompileScript.sml @@ -1,19 +1,15 @@ (* Compiles the CNF + PB checker *) -open preamble compilationLib cnfProgTheory +open preamble cnfProgTheory eval_cake_compile_x64Lib + eval_cake_compile_arm8Lib val _ = new_theory "cnfCompile" -val _ = intermediate_prog_prefix := "arm8_" +Theorem cnf_compiled = + eval_cake_compile_x64 "" main_prog_def "cake_pb_cnf.S"; -val cnf_compiled_arm8 = save_thm("cnf_compiled_arm8", - compile_arm8 "cake_pb_cnf_arm8" main_prog_def); - -(* Default has no prefix *) -val _ = intermediate_prog_prefix := "" - -val cnf_compiled = save_thm("cnf_compiled", - compile_x64 "cake_pb_cnf" main_prog_def); +Theorem cnf_compiled_arm8 = + eval_cake_compile_arm8 "arm8_" main_prog_def "cake_pb_cnf_arm8.S"; val _ = export_theory (); diff --git a/examples/pseudo_bool/array/compilation/mccisCompileScript.sml b/examples/pseudo_bool/array/compilation/mccisCompileScript.sml index f723fb839c..51ec069c56 100644 --- a/examples/pseudo_bool/array/compilation/mccisCompileScript.sml +++ b/examples/pseudo_bool/array/compilation/mccisCompileScript.sml @@ -1,19 +1,15 @@ (* Compiles the MCCIS + PB checker *) -open preamble compilationLib mccisProgTheory +open preamble mccisProgTheory eval_cake_compile_x64Lib + eval_cake_compile_arm8Lib val _ = new_theory "mccisCompile" -val _ = intermediate_prog_prefix := "arm8_" +Theorem mccis_compiled = + eval_cake_compile_x64 "" main_prog_def "cake_pb_mccis.S"; -val mccis_compiled_arm8 = save_thm("mccis_compiled_arm8", - compile_arm8 "cake_pb_mccis_arm8" main_prog_def); - -(* Default has no prefix *) -val _ = intermediate_prog_prefix := "" - -val mccis_compiled = save_thm("mccis_compiled", - compile_x64 "cake_pb_mccis" main_prog_def); +Theorem mccis_compiled_arm8 = + eval_cake_compile_arm8 "arm8_" main_prog_def "cake_pb_mccis_arm8.S"; val _ = export_theory (); diff --git a/examples/pseudo_bool/array/compilation/mcisCompileScript.sml b/examples/pseudo_bool/array/compilation/mcisCompileScript.sml index 4b85c82873..6119b95adb 100644 --- a/examples/pseudo_bool/array/compilation/mcisCompileScript.sml +++ b/examples/pseudo_bool/array/compilation/mcisCompileScript.sml @@ -1,19 +1,15 @@ (* Compiles the MCIS + PB checker *) -open preamble compilationLib mcisProgTheory +open preamble mcisProgTheory eval_cake_compile_x64Lib + eval_cake_compile_arm8Lib val _ = new_theory "mcisCompile" -val _ = intermediate_prog_prefix := "arm8_" +Theorem mcis_compiled = + eval_cake_compile_x64 "" main_prog_def "cake_pb_mcis.S"; -val mcis_compiled_arm8 = save_thm("mcis_compiled_arm8", - compile_arm8 "cake_pb_mcis_arm8" main_prog_def); - -(* Default has no prefix *) -val _ = intermediate_prog_prefix := "" - -val mcis_compiled = save_thm("mcis_compiled", - compile_x64 "cake_pb_mcis" main_prog_def); +Theorem mcis_compiled_arm8 = + eval_cake_compile_arm8 "arm8_" main_prog_def "cake_pb_mcis_arm8.S"; val _ = export_theory (); diff --git a/examples/pseudo_bool/array/compilation/npbc_fullCompileScript.sml b/examples/pseudo_bool/array/compilation/npbc_fullCompileScript.sml index 979b9870ea..8e2f2b7ffa 100644 --- a/examples/pseudo_bool/array/compilation/npbc_fullCompileScript.sml +++ b/examples/pseudo_bool/array/compilation/npbc_fullCompileScript.sml @@ -1,19 +1,15 @@ (* Compiles the PB checker example by evaluation inside the logic of HOL *) -open preamble compilationLib npbc_fullProgTheory +open preamble npbc_fullProgTheory eval_cake_compile_x64Lib + eval_cake_compile_arm8Lib val _ = new_theory "npbc_fullCompile" -val _ = intermediate_prog_prefix := "arm8_" +Theorem npbc_full_compiled = + eval_cake_compile_x64 "" main_prog_def "cake_pb.S"; -val pb_arrayFull_compiled_arm8 = save_thm("npbc_full_compiled_arm8", - compile_arm8 "cake_pb_arm8" main_prog_def); - -(* Default has no prefix *) -val _ = intermediate_prog_prefix := "" - -val pb_arrayFull_compiled = save_thm("npbc_full_compiled", - compile_x64 "cake_pb" main_prog_def); +Theorem npbc_full_compiled_arm8 = + eval_cake_compile_arm8 "arm8_" main_prog_def "cake_pb_arm8.S"; val _ = export_theory (); diff --git a/examples/pseudo_bool/array/compilation/proofs/cliqueProofScript.sml b/examples/pseudo_bool/array/compilation/proofs/cliqueProofScript.sml index 8e3745793b..433995518c 100644 --- a/examples/pseudo_bool/array/compilation/proofs/cliqueProofScript.sml +++ b/examples/pseudo_bool/array/compilation/proofs/cliqueProofScript.sml @@ -19,7 +19,7 @@ val (cake_pb_clique_sem,cake_pb_clique_output) = cake_pb_clique_io_events_def |> val (cake_pb_clique_not_fail,cake_pb_clique_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail cake_pb_clique_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct clique_compiled + MATCH_MP compile_correct (cj 1 clique_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP cake_pb_clique_not_fail |> C MATCH_MP x64_backend_config_ok @@ -52,7 +52,7 @@ val installed_x64_def = Define ` `; val cake_pb_clique_code_def = Define ` - cake_pb_clique_code = (code, data, config) + cake_pb_clique_code = (code, data, info) `; (* A standard run of cake_pb_clique diff --git a/examples/pseudo_bool/array/compilation/proofs/cnfProofScript.sml b/examples/pseudo_bool/array/compilation/proofs/cnfProofScript.sml index c5395edfb2..eceb810982 100644 --- a/examples/pseudo_bool/array/compilation/proofs/cnfProofScript.sml +++ b/examples/pseudo_bool/array/compilation/proofs/cnfProofScript.sml @@ -19,7 +19,7 @@ val (cake_pb_cnf_sem,cake_pb_cnf_output) = cake_pb_cnf_io_events_def |> SPEC_ALL val (cake_pb_cnf_not_fail,cake_pb_cnf_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail cake_pb_cnf_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct cnf_compiled + MATCH_MP compile_correct (cj 1 cnf_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP cake_pb_cnf_not_fail |> C MATCH_MP x64_backend_config_ok @@ -52,7 +52,7 @@ val installed_x64_def = Define ` `; val cake_pb_cnf_code_def = Define ` - cake_pb_cnf_code = (code, data, config) + cake_pb_cnf_code = (code, data, info) `; (* A standard run of cake_pb_cnf diff --git a/examples/pseudo_bool/array/compilation/proofs/mccisProofScript.sml b/examples/pseudo_bool/array/compilation/proofs/mccisProofScript.sml index 85fc376d5b..a92678f496 100644 --- a/examples/pseudo_bool/array/compilation/proofs/mccisProofScript.sml +++ b/examples/pseudo_bool/array/compilation/proofs/mccisProofScript.sml @@ -19,7 +19,7 @@ val (cake_pb_mccis_sem,cake_pb_mccis_output) = cake_pb_mccis_io_events_def |> SP val (cake_pb_mccis_not_fail,cake_pb_mccis_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail cake_pb_mccis_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct mccis_compiled + MATCH_MP compile_correct (cj 1 mccis_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP cake_pb_mccis_not_fail |> C MATCH_MP x64_backend_config_ok @@ -52,7 +52,7 @@ val installed_x64_def = Define ` `; val cake_pb_mccis_code_def = Define ` - cake_pb_mccis_code = (code, data, config) + cake_pb_mccis_code = (code, data, info) `; (* A standard run of cake_pb_mccis diff --git a/examples/pseudo_bool/array/compilation/proofs/mcisProofScript.sml b/examples/pseudo_bool/array/compilation/proofs/mcisProofScript.sml index b91bcac190..e2b9ffbe79 100644 --- a/examples/pseudo_bool/array/compilation/proofs/mcisProofScript.sml +++ b/examples/pseudo_bool/array/compilation/proofs/mcisProofScript.sml @@ -19,7 +19,7 @@ val (cake_pb_mcis_sem,cake_pb_mcis_output) = cake_pb_mcis_io_events_def |> SPEC_ val (cake_pb_mcis_not_fail,cake_pb_mcis_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail cake_pb_mcis_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct mcis_compiled + MATCH_MP compile_correct (cj 1 mcis_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP cake_pb_mcis_not_fail |> C MATCH_MP x64_backend_config_ok @@ -52,7 +52,7 @@ val installed_x64_def = Define ` `; val cake_pb_mcis_code_def = Define ` - cake_pb_mcis_code = (code, data, config) + cake_pb_mcis_code = (code, data, info) `; (* A standard run of cake_pb_mcis diff --git a/examples/pseudo_bool/array/compilation/proofs/npbc_fullProofScript.sml b/examples/pseudo_bool/array/compilation/proofs/npbc_fullProofScript.sml index f1180021ec..02df97b48e 100644 --- a/examples/pseudo_bool/array/compilation/proofs/npbc_fullProofScript.sml +++ b/examples/pseudo_bool/array/compilation/proofs/npbc_fullProofScript.sml @@ -19,7 +19,7 @@ val (cake_pb_sem,cake_pb_output) = cake_pb_io_events_def |> SPEC_ALL |> UNDISCH val (cake_pb_not_fail,cake_pb_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail cake_pb_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct npbc_full_compiled + MATCH_MP compile_correct (cj 1 npbc_full_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP cake_pb_not_fail |> C MATCH_MP x64_backend_config_ok @@ -52,7 +52,7 @@ val installed_x64_def = Define ` `; val cake_pb_code_def = Define ` - cake_pb_code = (code, data, config) + cake_pb_code = (code, data, info) `; (* A standard run of cake_pb diff --git a/examples/pseudo_bool/array/compilation/proofs/subgraph_isoProofScript.sml b/examples/pseudo_bool/array/compilation/proofs/subgraph_isoProofScript.sml index 3cb3fa9d54..77cf93e3fc 100644 --- a/examples/pseudo_bool/array/compilation/proofs/subgraph_isoProofScript.sml +++ b/examples/pseudo_bool/array/compilation/proofs/subgraph_isoProofScript.sml @@ -19,7 +19,7 @@ val (cake_pb_iso_sem,cake_pb_iso_output) = cake_pb_iso_io_events_def |> SPEC_ALL val (cake_pb_iso_not_fail,cake_pb_iso_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail cake_pb_iso_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct subgraph_iso_compiled + MATCH_MP compile_correct (cj 1 subgraph_iso_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP cake_pb_iso_not_fail |> C MATCH_MP x64_backend_config_ok @@ -52,7 +52,7 @@ val installed_x64_def = Define ` `; val cake_pb_iso_code_def = Define ` - cake_pb_iso_code = (code, data, config) + cake_pb_iso_code = (code, data, info) `; (* A standard run of cake_pb_subgraph_iso diff --git a/examples/pseudo_bool/array/compilation/proofs/wcnfProofScript.sml b/examples/pseudo_bool/array/compilation/proofs/wcnfProofScript.sml index d9e23003fb..d67521a8d9 100644 --- a/examples/pseudo_bool/array/compilation/proofs/wcnfProofScript.sml +++ b/examples/pseudo_bool/array/compilation/proofs/wcnfProofScript.sml @@ -19,7 +19,7 @@ val (cake_pb_wcnf_sem,cake_pb_wcnf_output) = cake_pb_wcnf_io_events_def |> SPEC_ val (cake_pb_wcnf_not_fail,cake_pb_wcnf_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail cake_pb_wcnf_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct wcnf_compiled + MATCH_MP compile_correct (cj 1 wcnf_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP cake_pb_wcnf_not_fail |> C MATCH_MP x64_backend_config_ok @@ -52,7 +52,7 @@ val installed_x64_def = Define ` `; val cake_pb_wcnf_code_def = Define ` - cake_pb_wcnf_code = (code, data, config) + cake_pb_wcnf_code = (code, data, info) `; (* A standard run of cake_pb_wcnf diff --git a/examples/pseudo_bool/array/compilation/proofsARM8/Holmakefile b/examples/pseudo_bool/array/compilation/proofsARM8/Holmakefile index 4ad9d618ea..ea2c5bc033 100644 --- a/examples/pseudo_bool/array/compilation/proofsARM8/Holmakefile +++ b/examples/pseudo_bool/array/compilation/proofsARM8/Holmakefile @@ -13,3 +13,9 @@ README_SOURCES = $(wildcard *Script.sml) $(wildcard *Lib.sml) $(wildcard *Syntax DIRS = $(wildcard */) README.md: $(CAKEMLDIR)/developers/readme_gen readmePrefix $(patsubst %,%readmePrefix,$(DIRS)) $(README_SOURCES) $(protect $(CAKEMLDIR)/developers/readme_gen) $(README_SOURCES) + +ifdef POLY + +HOLHEAP=$(HOLDIR)/examples/l3-machine-code/arm8/asl-equiv/l3-equivalence-heap + +endif diff --git a/examples/pseudo_bool/array/compilation/proofsARM8/npbc_fullProofARM8Script.sml b/examples/pseudo_bool/array/compilation/proofsARM8/npbc_fullProofARM8Script.sml index 3f48634549..cd50f1cb26 100644 --- a/examples/pseudo_bool/array/compilation/proofsARM8/npbc_fullProofARM8Script.sml +++ b/examples/pseudo_bool/array/compilation/proofsARM8/npbc_fullProofARM8Script.sml @@ -20,7 +20,7 @@ val (cake_pb_sem,cake_pb_output) = cake_pb_io_events_def |> SPEC_ALL |> UNDISCH val (cake_pb_not_fail,cake_pb_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail cake_pb_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct npbc_full_compiled_arm8 + MATCH_MP compile_correct (cj 1 npbc_full_compiled_arm8) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP cake_pb_not_fail |> C MATCH_MP arm8_configProofTheory.arm8_backend_config_ok @@ -53,7 +53,7 @@ val installed_arm8_asl_def = Define ` `; val cake_pb_code_def = Define ` - cake_pb_code = (arm8_code, arm8_data, arm8_config) + cake_pb_code = (arm8_code, arm8_data, arm8_info) `; (* A standard run of cake_pb diff --git a/examples/pseudo_bool/array/compilation/subgraph_isoCompileScript.sml b/examples/pseudo_bool/array/compilation/subgraph_isoCompileScript.sml index 58ef5386f7..b7598bd4c3 100644 --- a/examples/pseudo_bool/array/compilation/subgraph_isoCompileScript.sml +++ b/examples/pseudo_bool/array/compilation/subgraph_isoCompileScript.sml @@ -1,19 +1,15 @@ (* Compiles the encoder *) -open preamble compilationLib subgraph_isoProgTheory +open preamble subgraph_isoProgTheory eval_cake_compile_x64Lib + eval_cake_compile_arm8Lib val _ = new_theory "subgraph_isoCompile" -val _ = intermediate_prog_prefix := "arm8_" +Theorem subgraph_iso_compiled = + eval_cake_compile_x64 "" main_prog_def "cake_pb_iso.S"; -val subgraph_iso_compiled = save_thm("subgraph_iso_compiled_arm8", - compile_arm8 "cake_pb_iso_arm8" main_prog_def); - -(* Default has no prefix *) -val _ = intermediate_prog_prefix := "" - -val subgraph_iso_compiled = save_thm("subgraph_iso_compiled", - compile_x64 "cake_pb_iso" main_prog_def); +Theorem subgraph_iso_compiled_arm8 = + eval_cake_compile_arm8 "arm8_" main_prog_def "cake_pb_iso_arm8.S"; val _ = export_theory (); diff --git a/examples/pseudo_bool/array/compilation/wcnfCompileScript.sml b/examples/pseudo_bool/array/compilation/wcnfCompileScript.sml index 29bd33691c..44974ac80c 100644 --- a/examples/pseudo_bool/array/compilation/wcnfCompileScript.sml +++ b/examples/pseudo_bool/array/compilation/wcnfCompileScript.sml @@ -1,19 +1,15 @@ (* Compiles the WCNF + PB checker *) -open preamble compilationLib wcnfProgTheory +open preamble wcnfProgTheory eval_cake_compile_x64Lib + eval_cake_compile_arm8Lib val _ = new_theory "wcnfCompile" -val _ = intermediate_prog_prefix := "arm8_" +Theorem wcnf_compiled = + eval_cake_compile_x64 "" main_prog_def "cake_pb_wcnf.S"; -val wcnf_compiled_arm8 = save_thm("wcnf_compiled_arm8", - compile_arm8 "cake_pb_wcnf_arm8" main_prog_def); - -(* Default has no prefix *) -val _ = intermediate_prog_prefix := "" - -val wcnf_compiled = save_thm("wcnf_compiled", - compile_x64 "cake_pb_wcnf" main_prog_def); +Theorem wcnf_compiled_arm8 = + eval_cake_compile_arm8 "arm8_" main_prog_def "cake_pb_wcnf_arm8.S"; val _ = export_theory (); diff --git a/examples/sat_encodings/translation/compilation/Holmakefile b/examples/sat_encodings/translation/compilation/Holmakefile index e8b3a34e29..a000077a77 100644 --- a/examples/sat_encodings/translation/compilation/Holmakefile +++ b/examples/sat_encodings/translation/compilation/Holmakefile @@ -1,4 +1,4 @@ -INCLUDES = $(CAKEMLDIR)/examples/sat_encodings $(CAKEMLDIR)/examples/sat_encodings/translation $(CAKEMLDIR)/compiler +INCLUDES = $(CAKEMLDIR)/examples/sat_encodings $(CAKEMLDIR)/examples/sat_encodings/translation $(CAKEMLDIR)/compiler $(CAKEMLDIR)/cv_translator all: $(DEFAULT_TARGETS) README.md graphColoring_encoder killerSudoku_encoder nQueens_encoder numBoolRange_encoder sudoku_encoder .PHONY: all @@ -7,6 +7,10 @@ DIRS = $(wildcard */) README.md: $(CAKEMLDIR)/developers/readme_gen readmePrefix $(patsubst %,%readmePrefix,$(DIRS)) $(README_SOURCES) $(protect $(CAKEMLDIR)/developers/readme_gen) $(README_SOURCES) +ifdef POLY +HOLHEAP = $(CAKEMLDIR)/cv_translator/cake_compile_heap +endif + graphColoring_encoder: graphColoringEncoderCompileTheory.uo gcc -o graphColoring_encoder graphColoring_encoder.S $(CAKEMLDIR)/basis/basis_ffi.c diff --git a/examples/sat_encodings/translation/compilation/graphColoringEncoderCompileScript.sml b/examples/sat_encodings/translation/compilation/graphColoringEncoderCompileScript.sml index 97a2ee8d6e..64d34317d5 100644 --- a/examples/sat_encodings/translation/compilation/graphColoringEncoderCompileScript.sml +++ b/examples/sat_encodings/translation/compilation/graphColoringEncoderCompileScript.sml @@ -1,12 +1,12 @@ (* Compiles the encoder for the graph coloring problem *) - -open preamble compilationLib graphColoringEncoderProgTheory; +open preamble graphColoringEncoderProgTheory eval_cake_compile_x64Lib; val _ = new_theory "graphColoringEncoderCompile" Theorem graphColoring_encoder_compiled = - compile_x64 "graphColoring_encoder" graphColoring_encoder_prog_def; + eval_cake_compile_x64 "" graphColoring_encoder_prog_def + "graphColoring_encoder.S"; val _ = export_theory (); diff --git a/examples/sat_encodings/translation/compilation/killerSudokuEncoderCompileScript.sml b/examples/sat_encodings/translation/compilation/killerSudokuEncoderCompileScript.sml index efb7520c63..e3150ced37 100644 --- a/examples/sat_encodings/translation/compilation/killerSudokuEncoderCompileScript.sml +++ b/examples/sat_encodings/translation/compilation/killerSudokuEncoderCompileScript.sml @@ -1,12 +1,12 @@ (* Compile the encoder for the killer sudoku puzzle *) - -open preamble compilationLib killerSudokuEncoderProgTheory; +open preamble killerSudokuEncoderProgTheory eval_cake_compile_x64Lib; val _ = new_theory "killerSudokuEncoderCompile" Theorem killerSudoku_encoder_compiled = - compile_x64 "killerSudoku_encoder" killerSudoku_encoder_prog_def; + eval_cake_compile_x64 "" killerSudoku_encoder_prog_def + "killerSudoku_encoder.S"; val _ = export_theory (); diff --git a/examples/sat_encodings/translation/compilation/nQueensEncoderCompileScript.sml b/examples/sat_encodings/translation/compilation/nQueensEncoderCompileScript.sml index f6365000fe..001617550c 100644 --- a/examples/sat_encodings/translation/compilation/nQueensEncoderCompileScript.sml +++ b/examples/sat_encodings/translation/compilation/nQueensEncoderCompileScript.sml @@ -1,12 +1,12 @@ (* Compile the encoder for the n-queens problem *) - -open preamble compilationLib nQueensEncoderProgTheory; +open preamble nQueensEncoderProgTheory eval_cake_compile_x64Lib; val _ = new_theory "nQueensEncoderCompile" Theorem nQueens_encoder_compiled = - compile_x64 "nQueens_encoder" nQueens_encoder_prog_def; + eval_cake_compile_x64 "" nQueens_encoder_prog_def + "nQueens_encoder.S"; val _ = export_theory (); diff --git a/examples/sat_encodings/translation/compilation/numBoolRangeEncoderCompileScript.sml b/examples/sat_encodings/translation/compilation/numBoolRangeEncoderCompileScript.sml index 122df567ff..89b0192f76 100644 --- a/examples/sat_encodings/translation/compilation/numBoolRangeEncoderCompileScript.sml +++ b/examples/sat_encodings/translation/compilation/numBoolRangeEncoderCompileScript.sml @@ -1,12 +1,12 @@ (* Compile the encoder for the numBoolRange datatype *) - -open preamble compilationLib numBoolRangeEncoderProgTheory; +open preamble numBoolRangeEncoderProgTheory eval_cake_compile_x64Lib; val _ = new_theory "numBoolRangeEncoderCompile" Theorem numBoolRange_encoder_compiled = - compile_x64 "numBoolRange_encoder" numBoolRange_encoder_prog_def; + eval_cake_compile_x64 "" numBoolRange_encoder_prog_def + "numBoolRange_encoder.S"; val _ = export_theory (); diff --git a/examples/sat_encodings/translation/compilation/sudokuEncoderCompileScript.sml b/examples/sat_encodings/translation/compilation/sudokuEncoderCompileScript.sml index 3ca6dad8f3..93c8a3fe62 100644 --- a/examples/sat_encodings/translation/compilation/sudokuEncoderCompileScript.sml +++ b/examples/sat_encodings/translation/compilation/sudokuEncoderCompileScript.sml @@ -1,12 +1,12 @@ (* Compile the encoder for the sudoku puzzle *) - -open preamble compilationLib sudokuEncoderProgTheory; +open preamble sudokuEncoderProgTheory eval_cake_compile_x64Lib; val _ = new_theory "sudokuEncoderCompile" Theorem sudoku_encoder_compiled = - compile_x64 "sudoku_encoder" sudoku_encoder_prog_def; + eval_cake_compile_x64 "" sudoku_encoder_prog_def + "sudoku_encoder.S"; val _ = export_theory (); diff --git a/examples/vipr/compilation/Holmakefile b/examples/vipr/compilation/Holmakefile index c20bc6381a..fe45f1c842 100644 --- a/examples/vipr/compilation/Holmakefile +++ b/examples/vipr/compilation/Holmakefile @@ -1,4 +1,4 @@ -INCLUDES = $(CAKEMLDIR)/examples/vipr $(CAKEMLDIR)/compiler +INCLUDES = $(CAKEMLDIR)/examples/vipr $(CAKEMLDIR)/compiler $(CAKEMLDIR)/cv_translator all: $(DEFAULT_TARGETS) README.md cake_vipr .PHONY: all @@ -7,5 +7,9 @@ DIRS = $(wildcard */) README.md: $(CAKEMLDIR)/developers/readme_gen readmePrefix $(patsubst %,%readmePrefix,$(DIRS)) $(README_SOURCES) $(protect $(CAKEMLDIR)/developers/readme_gen) $(README_SOURCES) +ifdef POLY +HOLHEAP = $(CAKEMLDIR)/cv_translator/cake_compile_heap +endif + cake_vipr: viprCompileTheory.uo gcc -o cake_vipr cake_vipr.S $(CAKEMLDIR)/basis/basis_ffi.c diff --git a/examples/vipr/compilation/viprCompileScript.sml b/examples/vipr/compilation/viprCompileScript.sml index 645c8c4ea3..2aaf9ef5de 100644 --- a/examples/vipr/compilation/viprCompileScript.sml +++ b/examples/vipr/compilation/viprCompileScript.sml @@ -1,12 +1,11 @@ (* Compile cake_vipr *) - -open preamble compilationLib viprProgTheory; +open preamble viprProgTheory eval_cake_compile_x64Lib val _ = new_theory "viprCompile" Theorem vipr_encoder_compiled = - compile_x64 "cake_vipr" vipr_prog_def; + eval_cake_compile_x64 "" vipr_prog_def "cake_vipr.S"; val _ = export_theory (); diff --git a/examples/xlrup_checker/array/compilation/Holmakefile b/examples/xlrup_checker/array/compilation/Holmakefile index 8e7cea16c9..bef537dbe0 100644 --- a/examples/xlrup_checker/array/compilation/Holmakefile +++ b/examples/xlrup_checker/array/compilation/Holmakefile @@ -1,4 +1,4 @@ -INCLUDES = $(CAKEMLDIR)/misc $(CAKEMLDIR)/basis $(CAKEMLDIR)/compiler .. +INCLUDES = $(CAKEMLDIR)/misc $(CAKEMLDIR)/basis $(CAKEMLDIR)/compiler $(CAKEMLDIR)/cv_translator .. CLINE_OPTIONS = all: $(DEFAULT_TARGETS) README.md @@ -9,6 +9,10 @@ DIRS = $(wildcard */) README.md: $(CAKEMLDIR)/developers/readme_gen readmePrefix $(patsubst %,%readmePrefix,$(DIRS)) $(README_SOURCES) $(protect $(CAKEMLDIR)/developers/readme_gen) $(README_SOURCES) +ifdef POLY +HOLHEAP = $(CAKEMLDIR)/cv_translator/cake_compile_heap +endif + cake_lpr.S: *lpr_arrayCompileScript.sml cake_lpr_arm8.S: *lpr_arrayCompileARM8Script.sml cake_lpr_ramsey.S: *lpr_arrayRamseyCompileScript.sml diff --git a/examples/xlrup_checker/array/compilation/README.md b/examples/xlrup_checker/array/compilation/README.md index 18e7f32884..4f578a84be 100644 --- a/examples/xlrup_checker/array/compilation/README.md +++ b/examples/xlrup_checker/array/compilation/README.md @@ -1,4 +1,7 @@ An XLRUP checker built on CakeML with arrays +[proofs](proofs): +Prove end-to-end correctness theorem for xlrup checker with arrays + [xlrupCompileScript.sml](xlrupCompileScript.sml): Compiles the xlrup example by evaluation inside the logic of HOL diff --git a/examples/xlrup_checker/array/compilation/proofs/xlrupProofScript.sml b/examples/xlrup_checker/array/compilation/proofs/xlrupProofScript.sml index 3987fdf305..6616f58f14 100644 --- a/examples/xlrup_checker/array/compilation/proofs/xlrupProofScript.sml +++ b/examples/xlrup_checker/array/compilation/proofs/xlrupProofScript.sml @@ -19,7 +19,7 @@ val (cake_xlrup_sem,cake_xlrup_output) = cake_xlrup_io_events_def |> SPEC_ALL |> val (cake_xlrup_not_fail,cake_xlrup_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail cake_xlrup_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct xlrup_array_compiled + MATCH_MP compile_correct (cj 1 xlrup_array_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP cake_xlrup_not_fail |> C MATCH_MP x64_backend_config_ok @@ -52,7 +52,7 @@ val installed_x64_def = Define ` `; val cake_xlrup_code_def = Define ` - cake_xlrup_code = (code, data, config) + cake_xlrup_code = (code, data, info) `; (* A standard run of cake_xlrup satisfying all the default assumptions *) diff --git a/examples/xlrup_checker/array/compilation/xlrupCompileScript.sml b/examples/xlrup_checker/array/compilation/xlrupCompileScript.sml index 9db202256c..e44f1ae871 100644 --- a/examples/xlrup_checker/array/compilation/xlrupCompileScript.sml +++ b/examples/xlrup_checker/array/compilation/xlrupCompileScript.sml @@ -1,11 +1,11 @@ (* Compiles the xlrup example by evaluation inside the logic of HOL *) -open preamble compilationLib xlrup_arrayFullProgTheory +open preamble xlrup_arrayFullProgTheory eval_cake_compile_x64Lib val _ = new_theory "xlrupCompile" -val xlrup_array_compiled = save_thm("xlrup_array_compiled", - compile_x64 "cake_xlrup" check_unsat_prog_def); +Theorem xlrup_array_compiled = + eval_cake_compile_x64 "" check_unsat_prog_def "cake_xlrup.S"; val _ = export_theory (); diff --git a/misc/miscScript.sml b/misc/miscScript.sml index de0369c12b..1fa2d88579 100644 --- a/misc/miscScript.sml +++ b/misc/miscScript.sml @@ -213,6 +213,13 @@ Proof Cases \\ Cases \\ rw[] QED +Theorem SmartAppend_Nil: + SmartAppend l Nil = l ∧ + SmartAppend Nil l = l +Proof + Cases_on ‘l’ \\ gvs [SmartAppend_def] +QED + Theorem append_SmartAppend[simp]: append (SmartAppend l1 l2) = append l1 ++ append l2 Proof diff --git a/tutorial/simple_bstScript.sml b/tutorial/simple_bstScript.sml index c91272b8df..20c2585aef 100644 --- a/tutorial/simple_bstScript.sml +++ b/tutorial/simple_bstScript.sml @@ -25,14 +25,21 @@ open preamble comparisonTheory val _ = new_theory "simple_bst"; +val _ = Parse.hide "cmp"; + +Overload Less = “ternaryComparisons$LESS” +Overload Equal = “ternaryComparisons$EQUAL” +Overload Greater = “ternaryComparisons$GREATER” + (* Define the binary tree type. It is polymorphic with two type variables, 'k for the key type and 'v for the value type. *) -val _ = Datatype` - btree = Leaf | Node 'k 'v btree btree`; +Datatype: + btree = Leaf | Node 'k 'v btree btree +End (* Try, for example @@ -51,19 +58,22 @@ type_of ``Node``; TypeBase.constructors_of``:cpn``; *) -val singleton_def = Define` - singleton k v = Node k v Leaf Leaf`; +Definition singleton_def: + singleton k v = Node k v Leaf Leaf +End -val lookup_def = Define` +Definition lookup_def: lookup cmp k Leaf = NONE ∧ lookup cmp k (Node k' v' l r) = case cmp k k' of | Less => lookup cmp k l | Greater => lookup cmp k r - | Equal => SOME v'`; + | Equal => SOME v' +End -val insert_def = +Definition insert_def: (* EXERCISE: fill in your definition here *) +End (* Since we are working with an abstract comparison function, different keys (k, @@ -74,8 +84,9 @@ val insert_def = which reveals that this is defined in comparisonTheory *) -val key_set_def = Define` - key_set cmp k = { k' | cmp k k' = Equal } `; +Definition key_set_def: + key_set cmp k = { k' | cmp k k' = Equal } +End (* For a good comparison function cmp, the key_set cmp relation should be an @@ -113,8 +124,8 @@ QED (* A helper theorem, expanding out the definition of key_set, for use with metis_tac later. *) -val IN_key_set = save_thm("IN_key_set", - ``k' ∈ key_set cmp k`` |> SIMP_CONV (srw_ss()) [key_set_def]); +Theorem IN_key_set = + “k' ∈ key_set cmp k” |> SIMP_CONV (srw_ss()) [key_set_def]; (* Now let us define the (abstract) finite map from key-equivalence-classes to @@ -122,10 +133,11 @@ val IN_key_set = save_thm("IN_key_set", standard against which we can verify correctness of the operations. *) -val to_fmap_def = Define` +Definition to_fmap_def: to_fmap cmp Leaf = FEMPTY ∧ to_fmap cmp (Node k v l r) = - to_fmap cmp l ⊌ to_fmap cmp r |+ (key_set cmp k, v)`; + to_fmap cmp l ⊌ to_fmap cmp r |+ (key_set cmp k, v) +End Theorem to_fmap_key_set: ∀ks t. @@ -144,24 +156,26 @@ QED We start by defining what a predicate on trees indicating whether they have the binary search tree property *) -val key_ordered_def = Define` +Definition key_ordered_def: (key_ordered cmp k Leaf res ⇔ T) ∧ (key_ordered cmp k (Node k' _ l r) res ⇔ cmp k k' = res ∧ key_ordered cmp k l res ∧ - key_ordered cmp k r res)`; + key_ordered cmp k r res) +End (* We make this definition an automatic rewrite, so it is expanded automatically by simplification tactics (such as rw, fs, and simp) *) val _ = export_rewrites["key_ordered_def"]; -val wf_tree_def = Define` +Definition wf_tree_def: (wf_tree cmp Leaf ⇔ T) ∧ (wf_tree cmp (Node k _ l r) ⇔ key_ordered cmp k l Greater ∧ key_ordered cmp k r Less ∧ wf_tree cmp l ∧ - wf_tree cmp r)`; + wf_tree cmp r) +End (* We can prove that all the operations preserve wf_tree diff --git a/tutorial/solutions/Holmakefile b/tutorial/solutions/Holmakefile index 46361738c9..ae8abf73f3 100644 --- a/tutorial/solutions/Holmakefile +++ b/tutorial/solutions/Holmakefile @@ -2,7 +2,8 @@ INCLUDES = $(HOLDIR)/examples/balanced_bst\ $(CAKEMLDIR)/misc $(CAKEMLDIR)/semantics $(CAKEMLDIR)/semantics/proofs\ $(CAKEMLDIR)/basis/pure $(CAKEMLDIR)/basis\ $(CAKEMLDIR)/translator $(CAKEMLDIR)/characteristic\ - $(CAKEMLDIR)/compiler $(CAKEMLDIR)/compiler/backend/proofs\ + $(CAKEMLDIR)/cv_translator\ + $(CAKEMLDIR)/compiler $(CAKEMLDIR)/compiler/backend/proofs\ $(CAKEMLDIR)/compiler/backend/x64/proofs all: $(DEFAULT_TARGETS) README.md exercises @@ -15,6 +16,10 @@ README.md: $(CAKEMLDIR)/developers/readme_gen readmePrefix $(patsubst %,%readmeP SOLUTIONS = wordfreqProgScript.sml simple_bstScript.sml +ifdef POLY +HOLHEAP = $(CAKEMLDIR)/cv_translator/cake_compile_heap +endif + exercises: $(SOLUTIONS) make_ex ./make_ex $(SOLUTIONS) diff --git a/tutorial/solutions/simple_bstScript.sml b/tutorial/solutions/simple_bstScript.sml index d2ff971532..1ebbc64a27 100644 --- a/tutorial/solutions/simple_bstScript.sml +++ b/tutorial/solutions/simple_bstScript.sml @@ -25,14 +25,21 @@ open preamble comparisonTheory val _ = new_theory "simple_bst"; +val _ = Parse.hide "cmp"; + +Overload Less = “ternaryComparisons$LESS” +Overload Equal = “ternaryComparisons$EQUAL” +Overload Greater = “ternaryComparisons$GREATER” + (* Define the binary tree type. It is polymorphic with two type variables, 'k for the key type and 'v for the value type. *) -val _ = Datatype` - btree = Leaf | Node 'k 'v btree btree`; +Datatype: + btree = Leaf | Node 'k 'v btree btree +End (* Try, for example @@ -51,28 +58,30 @@ type_of ``Node``; TypeBase.constructors_of``:cpn``; *) -val singleton_def = Define` - singleton k v = Node k v Leaf Leaf`; +Definition singleton_def: + singleton k v = Node k v Leaf Leaf +End -val lookup_def = Define` +Definition lookup_def: lookup cmp k Leaf = NONE ∧ lookup cmp k (Node k' v' l r) = case cmp k k' of | Less => lookup cmp k l | Greater => lookup cmp k r - | Equal => SOME v'`; + | Equal => SOME v' +End -val insert_def = +Definition insert_def: (* EXERCISE: fill in your definition here *) (*ex *) -Define` insert cmp k v Leaf = singleton k v ∧ insert cmp k v (Node k' v' l r) = case cmp k k' of | Less => Node k' v' (insert cmp k v l) r | Greater => Node k' v' l (insert cmp k v r) - | Equal => Node k' v l r`; + | Equal => Node k' v l r (* ex*) +End (* Since we are working with an abstract comparison function, different keys (k, @@ -83,8 +92,9 @@ Define` which reveals that this is defined in comparisonTheory *) -val key_set_def = Define` - key_set cmp k = { k' | cmp k k' = Equal } `; +Definition key_set_def: + key_set cmp k = { k' | cmp k k' = Equal } +End (* For a good comparison function cmp, the key_set cmp relation should be an @@ -126,8 +136,8 @@ QED (* A helper theorem, expanding out the definition of key_set, for use with metis_tac later. *) -val IN_key_set = save_thm("IN_key_set", - ``k' ∈ key_set cmp k`` |> SIMP_CONV (srw_ss()) [key_set_def]); +Theorem IN_key_set = + “k' ∈ key_set cmp k” |> SIMP_CONV (srw_ss()) [key_set_def]; (* Now let us define the (abstract) finite map from key-equivalence-classes to @@ -135,10 +145,11 @@ val IN_key_set = save_thm("IN_key_set", standard against which we can verify correctness of the operations. *) -val to_fmap_def = Define` +Definition to_fmap_def: to_fmap cmp Leaf = FEMPTY ∧ to_fmap cmp (Node k v l r) = - to_fmap cmp l ⊌ to_fmap cmp r |+ (key_set cmp k, v)`; + to_fmap cmp l ⊌ to_fmap cmp r |+ (key_set cmp k, v) +End Theorem to_fmap_key_set: ∀ks t. @@ -161,24 +172,26 @@ QED We start by defining what a predicate on trees indicating whether they have the binary search tree property *) -val key_ordered_def = Define` +Definition key_ordered_def: (key_ordered cmp k Leaf res ⇔ T) ∧ (key_ordered cmp k (Node k' _ l r) res ⇔ cmp k k' = res ∧ key_ordered cmp k l res ∧ - key_ordered cmp k r res)`; + key_ordered cmp k r res) +End (* We make this definition an automatic rewrite, so it is expanded automatically by simplification tactics (such as rw, fs, and simp) *) val _ = export_rewrites["key_ordered_def"]; -val wf_tree_def = Define` +Definition wf_tree_def: (wf_tree cmp Leaf ⇔ T) ∧ (wf_tree cmp (Node k _ l r) ⇔ key_ordered cmp k l Greater ∧ key_ordered cmp k r Less ∧ wf_tree cmp l ∧ - wf_tree cmp r)`; + wf_tree cmp r) +End (* We can prove that all the operations preserve wf_tree diff --git a/tutorial/wordcountCompileScript.sml b/tutorial/wordcountCompileScript.sml index fa970b6f4c..cd4deb5c86 100644 --- a/tutorial/wordcountCompileScript.sml +++ b/tutorial/wordcountCompileScript.sml @@ -3,11 +3,11 @@ in the logic. *) -open preamble wordcountProgTheory compilationLib +open preamble wordcountProgTheory eval_cake_compile_x64Lib val _ = new_theory"wordcountCompile"; -val wordcount_compiled = save_thm("wordcount_compiled", - compile_x64 "wordcount" wordcount_prog_def); +Theorem wordcount_compiled = + eval_cake_compile_x64 "" wordcount_prog_def "wordcount.S"; val _ = export_theory(); diff --git a/tutorial/wordcountProofScript.sml b/tutorial/wordcountProofScript.sml index 1d95fa89f2..a4e2efe016 100644 --- a/tutorial/wordcountProofScript.sml +++ b/tutorial/wordcountProofScript.sml @@ -19,7 +19,7 @@ val (wordcount_sem,wordcount_output) = wordcount_io_events_def |> SPEC_ALL |> UN val (wordcount_not_fail,wordcount_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail wordcount_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct wordcount_compiled + MATCH_MP compile_correct (cj 1 wordcount_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP wordcount_not_fail |> C MATCH_MP x64_backend_config_ok diff --git a/tutorial/wordfreqCompileScript.sml b/tutorial/wordfreqCompileScript.sml index 0208262953..683f8c1d43 100644 --- a/tutorial/wordfreqCompileScript.sml +++ b/tutorial/wordfreqCompileScript.sml @@ -2,12 +2,11 @@ Compile the wordfreq program to machine code by evaluation of the compiler in the logic. *) - -open preamble wordfreqProgTheory compilationLib +open preamble wordfreqProgTheory eval_cake_compile_x64Lib val _ = new_theory"wordfreqCompile"; -val wordfreq_compiled = save_thm("wordfreq_compiled", - compile_x64 "wordfreq" wordfreq_prog_def); +Theorem wordfreq_compiled = + eval_cake_compile_x64 "" wordfreq_prog_def "wordfreq.S"; val _ = export_theory(); diff --git a/tutorial/wordfreqProofScript.sml b/tutorial/wordfreqProofScript.sml index ca1f287c54..1854ec6501 100644 --- a/tutorial/wordfreqProofScript.sml +++ b/tutorial/wordfreqProofScript.sml @@ -17,7 +17,7 @@ val (wordfreq_sem,wordfreq_output) = wordfreq_io_events_def |> SPEC_ALL |> UNDIS val (wordfreq_not_fail,wordfreq_sem_sing) = MATCH_MP semantics_prog_Terminate_not_Fail wordfreq_sem |> CONJ_PAIR val compile_correct_applied = - MATCH_MP compile_correct wordfreq_compiled + MATCH_MP compile_correct (cj 1 wordfreq_compiled) |> SIMP_RULE(srw_ss())[LET_THM,ml_progTheory.init_state_env_thm,GSYM AND_IMP_INTRO] |> C MATCH_MP wordfreq_not_fail |> C MATCH_MP x64_backend_config_ok @@ -32,12 +32,13 @@ val wordfreq_compiled_lemma = |> SIMP_RULE std_ss [AND_IMP_INTRO,GSYM CONJ_ASSOC] |> DISCH_ALL; -val compiler_output_def = Define ` - compiler_output = (code,data,config)`; +Definition compiler_output_def: + compiler_output = (code,data,info) +End (* TODO: move *) -val get_file_contents_def = Define ` +Definition get_file_contents_def: get_file_contents fs fname = if inFS_fname fs fname then case ALOOKUP fs.files fname of @@ -46,16 +47,19 @@ val get_file_contents_def = Define ` case ALOOKUP fs.inode_tbl (File ino) of | NONE => NONE | SOME s => SOME (implode s) - else NONE` + else NONE +End -val wfFS_def = Define ` - wfFS fs <=> fsFFIProps$wfFS fs ∧ STD_streams fs`; +Definition wfFS_def: + wfFS fs <=> fsFFIProps$wfFS fs ∧ STD_streams fs +End -val x64_installed_def = Define ` +Definition x64_installed_def: x64_installed (c,d,conf) cbspace data_sp mc ms <=> is_x64_machine_config mc ∧ targetSem$installed c cbspace d data_sp conf.lab_conf.ffi_names - (heap_regs x64_backend_config.stack_conf.reg_names) mc conf.lab_conf.shmem_extra ms` + (heap_regs x64_backend_config.stack_conf.reg_names) mc conf.lab_conf.shmem_extra ms +End (* -- *) @@ -71,7 +75,7 @@ Theorem wordfreq_compiled_thm: Proof strip_tac \\ assume_tac wordfreq_compiled_lemma - \\ rfs [get_file_contents_def,wfFS_def,compiler_output_def,x64_installed_def] + \\ rfs [get_file_contents_def,wfFS_def,x64_installed_def,compiler_output_def] \\ asm_exists_tac \\ fs [option_case_eq] \\ metis_tac [wordfreqProgTheory.wordfreq_output_spec_def] QED