From dfe6245e933c8b1a4d08d7bf9c3822628f04a597 Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Thu, 18 Jan 2024 10:58:13 +0800 Subject: [PATCH 01/38] de-duplicate in wcnf_to_pb, better errors --- .../array/npbc_arrayProgScript.sml | 152 +++++++++++------- .../cnf_encoding/wcnf_to_pbScript.sml | 38 +++-- 2 files changed, 123 insertions(+), 67 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_arrayProgScript.sml b/examples/pseudo_bool/array/npbc_arrayProgScript.sml index 1259a32f0b..7f0190bc77 100644 --- a/examples/pseudo_bool/array/npbc_arrayProgScript.sml +++ b/examples/pseudo_bool/array/npbc_arrayProgScript.sml @@ -1481,10 +1481,11 @@ val r = translate splim_def; val every_hs = process_topdecs` fun every_hs hs ls = - case ls of [] => True + case ls of [] => None | l::ls => - in_hashset_arr l hs andalso - every_hs hs ls` |> append_prog + if in_hashset_arr l hs then + every_hs hs ls + else Some (npbc_string l)` |> append_prog Theorem every_hs_spec: ∀ls lsv. @@ -1498,26 +1499,30 @@ Theorem every_hs_spec: (ARRAY hspv hsv) (POSTv resv. ARRAY hspv hsv * - &BOOL (EVERY (λc. in_hashset c hs) ls) resv) + & ∃err. + OPTION_TYPE STRING_TYPE + (if EVERY (λc. in_hashset c hs) ls + then NONE else SOME err) resv) Proof Induct>>rw[]>> fs[LIST_TYPE_def]>> xcf "every_hs" (get_ml_prog_state ())>> xmatch - >- (xcon>>xsimpl)>> + >- (xcon>>xsimpl>>EVAL_TAC)>> xlet_auto>> - xlog>> - xsimpl>>rw[] - >- ( - xapp>>xsimpl>> - pairarg_tac>>fs[])>> - gvs[] + xif + >- + (xapp>>xsimpl)>> + xlet_autop>> + xcon>>xsimpl>> + simp[OPTION_TYPE_def]>> + metis_tac[] QED val hash_check = process_topdecs` fun hash_check fmlls proved lf = case lf of - [] => True + [] => None | _ => let val hs = Array.array splim [] @@ -1544,10 +1549,12 @@ Theorem hash_check_spec: [xv; yv; lsv] emp (POSTv resv. - &BOOL ( - let hs = mk_hashset x + & ∃err. + OPTION_TYPE STRING_TYPE + (let hs = mk_hashset x (mk_hashset y (REPLICATE splim [])) in - EVERY (λc. in_hashset c hs) ls) resv) + if EVERY (λc. in_hashset c hs) ls + then NONE else SOME err) resv) Proof rw[]>> xcf "hash_check" (get_ml_prog_state ())>> @@ -1555,7 +1562,8 @@ Proof >- ( fs[LIST_TYPE_def]>> xmatch>> - xcon>>xsimpl)>> + xcon>>xsimpl>> + EVAL_TAC)>> fs[LIST_TYPE_def]>> xmatch>> xlet_autop>> @@ -1631,7 +1639,7 @@ val res = translate red_cond_check_pure_def; val red_cond_check = process_topdecs` fun red_cond_check indfml extra pfs rsubs goals = case red_cond_check_pure extra pfs rsubs goals of - None => False + None => Some "not all # subgoals present" | Some (x,ls) => hash_check indfml x ls` |> append_prog @@ -1649,18 +1657,22 @@ Theorem red_cond_check_spec: [av; aav; bv; cv; dv] emp (POSTv resv. - &BOOL (red_cond_check a aa b c d) resv) + & ∃err. + OPTION_TYPE STRING_TYPE + (if red_cond_check a aa b c d then NONE + else SOME err) resv) Proof rw[]>> xcf "red_cond_check" (get_ml_prog_state ())>> fs[]>> xlet_autop>> simp[red_cond_check_eq]>> - TOP_CASE_TAC>>fs[OPTION_TYPE_def] + Cases_on`red_cond_check_pure aa b c d`>> + fs[OPTION_TYPE_def] >- ( xmatch>> xcon>>xsimpl)>> - TOP_CASE_TAC>>fs[PAIR_TYPE_def]>> + Cases_on`x`>>fs[PAIR_TYPE_def]>> xmatch>> xapp>> xsimpl>> @@ -1749,10 +1761,10 @@ val check_red_arr = process_topdecs` None => let val u = rollback_arr fml' id id' val goals = subst_indexes_arr s bortcb fml' rinds in - if red_cond_check fmlls nc pfs rsubs goals - then - (fml', (rinds,id')) - else raise Fail (format_failure_2 lno ("Redundancy subproofs did not cover all subgoals.") (print_subproofs_err rsubs goals)) + case red_cond_check fmlls nc pfs rsubs goals + of None => (fml', (rinds,id')) + | Some err => + raise Fail (format_failure_2 lno ("Redundancy subproofs did not cover all subgoals. Info: " ^ err ^ ".") (print_subproofs_err rsubs goals)) end | Some cid => if check_contradiction_fml_arr b fml' cid then @@ -1950,20 +1962,23 @@ Proof >- ( rpt xlet_autop>> fs[do_red_check_def]>> - reverse xif + pop_assum mp_tac>>IF_CASES_TAC>> + strip_tac>>fs[OPTION_TYPE_def]>>xmatch + >- ( + fs[red_cond_check_def]>> + pairarg_tac>>fs[]>> + xlet_autop>> + xcon>>xsimpl>> + simp[PAIR_TYPE_def]>> + xsimpl) >- ( rpt xlet_autop>> fs[red_cond_check_def]>> xraise>> xsimpl>> rw[]>>fs[]>> - metis_tac[ARRAY_refl,NOT_EVERY,Fail_exn_def]) >> - fs[red_cond_check_def]>> - pairarg_tac>>fs[]>> - xlet_autop>> - xcon>>xsimpl>> - simp[PAIR_TYPE_def]>> - xsimpl)>> + metis_tac[ARRAY_refl,NOT_EVERY,Fail_exn_def]) + )>> rpt xlet_autop>> reverse xif >- ( @@ -2303,10 +2318,10 @@ val check_dom_arr = process_topdecs` None => let val u = rollback_arr fml' id id' val goals = core_subgoals s corels in - if red_cond_check fmlls nc pfs dsubs goals - then - (fml',(rinds,id')) - else raise Fail (format_failure_2 lno ("Dominance subproofs did not cover all subgoals") (print_subproofs_err dsubs goals)) + case red_cond_check fmlls nc pfs dsubs goals of + None => (fml',(rinds,id')) + | Some err => + raise Fail (format_failure_2 lno ("Dominance subproofs did not cover all subgoals. Info: " ^ err ^ ".") (print_subproofs_err dsubs goals)) end | Some cid => if check_contradiction_fml_arr False fml' cid then @@ -2446,20 +2461,23 @@ Proof xmatch >- ( rpt xlet_autop>> - reverse xif>>gs[] + pop_assum mp_tac>>IF_CASES_TAC>> + strip_tac>>fs[OPTION_TYPE_def]>>xmatch + >- ( + xlet_autop>> + xcon>>xsimpl>> + pairarg_tac>>fs[red_cond_check_def,core_subgoals_def]>> + fs[PAIR_TYPE_def,OPTION_TYPE_def]>> + xsimpl) >- ( rpt xlet_autop>> fs[red_cond_check_def]>>pairarg_tac>>fs[core_subgoals_def]>> xraise>>xsimpl>> fs[red_cond_check_def]>>rw[]>> - metis_tac[ARRAY_refl,Fail_exn_def,NOT_EVERY])>> - xlet_autop>> - xcon>>xsimpl>> - pairarg_tac>>fs[red_cond_check_def,core_subgoals_def]>> - fs[PAIR_TYPE_def,OPTION_TYPE_def]>> - xsimpl)>> - rpt xlet_autop>> + metis_tac[ARRAY_refl,Fail_exn_def,NOT_EVERY]) + )>> rename1`check_contradiction_fml_list F A B`>> + xlet_autop>> xlet`POSTv v. ARRAY fmlv' fmllsv' * & BOOL (check_contradiction_fml_list F A B) v` @@ -3616,7 +3634,10 @@ Theorem fml_include_arr_spec: [lsv; rsv] (emp) (POSTv v. - &BOOL (fml_include_list ls rs) v) + & ∃err. + OPTION_TYPE STRING_TYPE + (if fml_include_list ls rs + then NONE else SOME err) v) Proof rw[fml_include_list_def]>> xcf"fml_include_arr"(get_ml_prog_state ())>> @@ -3636,24 +3657,33 @@ QED val _ = register_type ``:output`` +Definition print_opt_string_def: + (print_opt_string NONE = strlit"T") ∧ + (print_opt_string (SOME s) = s) +End + +val res = translate print_opt_string_def; + Definition derivable_res_def: - derivable_res (dbound :int option) b2 = + derivable_res (dbound :int option) b2opt = let b1 = (dbound = NONE) in + let b2 = (b2opt = NONE) in if b1 ∧ b2 then NONE else SOME ( strlit "output DERIVABLE check failed:" ^ strlit " [dbound = NONE] " ^ bool_to_string b1 ^ - strlit " [core in output] " ^ bool_to_string b2 - ) + strlit " [core in output] " ^ print_opt_string b2opt) End val res = translate derivable_res_def; Definition equisatisfiable_res_def: equisatisfiable_res - (bound:int option) (dbound:int option) chk b2 b3 = + (bound:int option) (dbound:int option) chk b2opt b3opt = let b11 = (bound = NONE) in let b12 = (dbound = NONE) in + let b2 = (b2opt = NONE) in + let b3 = (b3opt = NONE) in if b11 ∧ b12 ∧ chk ∧ b2 ∧ b3 then NONE else SOME ( @@ -3661,8 +3691,8 @@ Definition equisatisfiable_res_def: strlit " [bound = NONE] " ^ bool_to_string b11 ^ strlit " [dbound = NONE] " ^ bool_to_string b12 ^ strlit " [checked deletion] " ^ bool_to_string chk ^ - strlit " [core in output] " ^ bool_to_string b2 ^ - strlit " [output in core] " ^ bool_to_string b3 + strlit " [core in output] " ^ print_opt_string b2opt ^ + strlit " [output in core] " ^ print_opt_string b3opt ) End @@ -3670,21 +3700,33 @@ val res = translate equisatisfiable_res_def; val res = translate npbc_checkTheory.opt_eq_obj_def; +Definition opt_err_obj_check_string_def: + (opt_err_obj_check_string (SOME fc) (SOME fc') = + err_obj_check_string (SOME fc) fc') ∧ + (opt_err_obj_check_string _ _ = strlit "missing objective") +End + +val res = translate opt_err_obj_check_string_def; + Definition equioptimal_res_def: equioptimal_res (bound:int option) (dbound:int option) chk - obj obj' b2 b3 = + obj obj' b2opt b3opt = let b11 = opt_le bound dbound in let b12 = opt_eq_obj obj obj' in + let b12s = + if b12 then strlit"T" else opt_err_obj_check_string obj obj' in + let b2 = (b2opt = NONE) in + let b3 = (b3opt = NONE) in if b11 ∧ b12 ∧ chk ∧ b2 ∧ b3 then NONE else SOME ( strlit "output EQUIOPTIMAL check failed:" ^ strlit " [bound <= dbound]: " ^ bool_to_string b11 ^ - strlit " [obj = output obj]: " ^ bool_to_string b12 ^ + strlit " [obj = output obj]: " ^ b12s ^ strlit " [checked deletion]: " ^ bool_to_string chk ^ - strlit " [core in output]: " ^ bool_to_string b2 ^ - strlit " [output in core]: " ^ bool_to_string b3 + strlit " [core in output] " ^ print_opt_string b2opt ^ + strlit " [output in core] " ^ print_opt_string b3opt ) End diff --git a/examples/pseudo_bool/cnf_encoding/wcnf_to_pbScript.sml b/examples/pseudo_bool/cnf_encoding/wcnf_to_pbScript.sml index 26dadf3e50..2c3efdbc75 100644 --- a/examples/pseudo_bool/cnf_encoding/wcnf_to_pbScript.sml +++ b/examples/pseudo_bool/cnf_encoding/wcnf_to_pbScript.sml @@ -83,7 +83,7 @@ End ≤ 1 PB constraints and ≤ 1 terms in the objective *) Definition wclause_to_pbc_def: wclause_to_pbc (i,n,C) = - let C = FILTER (λl. l ≠ 0) C in + let C = nub (FILTER (λl. l ≠ 0) C) in if n = 0 then (* hard clauses *) ([(GreaterEqual,enc_clause C,1:int)],[]) else (* soft clauses *) @@ -239,6 +239,19 @@ Proof metis_tac[interp_cclause_FILTER] QED +Theorem weight_clause_nub: + weight_clause f (q,nub C) = + weight_clause f (q,C) +Proof + rw[weight_clause_def,interp_cclause_def] +QED + +Theorem interp_cclause_nub: + interp_cclause (nub C) = interp_cclause C +Proof + rw[interp_cclause_def] +QED + (* The sum of weights for unsatisfied clauses is upper bounded by the (negated) obj *) Theorem weight_clause_obj_upper: @@ -264,7 +277,7 @@ Proof fs[Abbr`C`,wf_clause_def,MEM_FILTER]>> `weight_clause (λx. w (INL x)) (q,r) = weight_clause (λx. w (INL x)) (q,C)` by - metis_tac[Abbr`C`,weight_clause_FILTER]>> + metis_tac[Abbr`C`,weight_clause_FILTER,weight_clause_nub]>> rw[]>>simp[weight_clause_def,iSUM_def] >- ( Cases_on`C`>>fs[]>> @@ -307,6 +320,7 @@ Proof Cases_on`EL n wfml`>> fs[wclause_to_pbc_def]>> strip_tac>>simp[Once interp_cclause_FILTER]>> + PURE_ONCE_REWRITE_TAC[GSYM interp_cclause_nub]>> match_mp_tac satisfies_pbc_satisfies_clause>> simp[wf_clause_def,MEM_FILTER])>> drule_all weight_clause_obj_upper>> @@ -359,14 +373,15 @@ Proof rw[] >- ( match_mp_tac satisfies_clause_satisfies_pbc>> - simp[GSYM interp_cclause_FILTER]>> - first_x_assum drule>>simp[wf_clause_def,MEM_FILTER])>> + simp[interp_cclause_nub,GSYM interp_cclause_FILTER]>> + first_x_assum drule>> + simp[wf_clause_def,MEM_FILTER])>> (* For the non-hard clauses *) Cases_on`satisfies_clause w (interp_cclause r)`>> simp[satisfies_pbc_def,eval_lin_term_def,iSUM_def]>> simp[GSYM eval_lin_term_def] >- ( - fs[Once interp_cclause_FILTER]>> + fs[Once interp_cclause_FILTER,Once (GSYM interp_cclause_nub)]>> drule_at Any satisfies_clause_satisfies_pbc>> simp[satisfies_pbc_def]>> disch_then match_mp_tac>> @@ -396,15 +411,14 @@ Proof `wf_clause C` by fs[Abbr`C`,wf_clause_def,MEM_FILTER]>> rw[]>>fs[iSUM_def]>> + `satisfies_clause w (interp_cclause r) = + satisfies_clause w (interp_cclause C)` by + metis_tac[Abbr`C`,interp_cclause_nub,interp_cclause_FILTER]>> + gvs[]>> ( Cases_on`C`>>fs[]>> - fs[Once interp_cclause_FILTER]>> gvs[interp_cclause_def,wf_clause_def,interp_lit_def,enc_lit_def]>> - rw[]>>fs[satisfies_clause_def,satisfies_literal_def]>> - IF_CASES_TAC>> - gvs[interp_cclause_def,satisfies_clause_def,wf_clause_def]>> - rw[enc_lit_def]>>fs[interp_lit_def,satisfies_literal_def]>> - intLib.ARITH_TAC))>> + rw[]>>fs[satisfies_clause_def,satisfies_literal_def]))>> pop_assum (qspec_then`k+1` sym_sub_tac)>> AP_TERM_TAC>> rw[MAP_EQ_f,MEM_FLAT,MEM_MAP,PULL_EXISTS]>> @@ -679,7 +693,7 @@ End strlit"h 1 2 3 4 0"; strlit"1 -3 -5 6 7 0"; strlit"6 -1 -2 0"; - strlit"4 1 6 -7 0";]`` + strlit"4 1 6 -7 6 -7 0";]`` val enc = EVAL`` full_encode (THE ^(rconc wcnf))`` *) From 75bbaa11104501a14549ee436d346282ff530467 Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sun, 21 Jan 2024 01:31:12 +0800 Subject: [PATCH 02/38] sketch tweaks to add earliest --- .../pseudo_bool/array/npbc_listScript.sml | 234 +++++++++++++++--- examples/pseudo_bool/npbc_checkScript.sml | 43 ++-- examples/pseudo_bool/pb_parseScript.sml | 34 +-- 3 files changed, 241 insertions(+), 70 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index cd4b7dd001..243a7355ff 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -288,13 +288,13 @@ Theorem check_lstep_list_id_del: SOME (fmlls', res) ∧ n < id ∧ IS_SOME (any_el n fmlls' NONE) ⇒ - IS_SOME (any_el n fmlls NONE)) ∧ + any_el n fmlls NONE = any_el n fmlls' NONE) ∧ (∀steps b fmlls mindel id fmlls' res n. check_lsteps_list steps b fmlls mindel id = SOME (fmlls', res) ∧ n < id ∧ IS_SOME (any_el n fmlls' NONE) ⇒ - IS_SOME (any_el n fmlls NONE)) + any_el n fmlls NONE = any_el n fmlls' NONE) Proof ho_match_mp_tac check_lstep_list_ind>> rw[] @@ -558,6 +558,27 @@ Definition reindex_def: (reindex b fml is = reindex_aux b fml is [] []) End +Definition reindex_partial_aux_def: + (reindex_partial_aux b fml mini [] iacc vacc = + (REVERSE iacc, vacc,[])) ∧ + (reindex_partial_aux b fml mini (i::is) iacc vacc = + if i < mini then (REVERSE iacc, vacc, i::is) + else + case any_el i fml NONE of + NONE => reindex_partial_aux b fml mini is iacc vacc + | SOME (v,b') => + let vacc' = + if b ⇒ b' then v::vacc else vacc in + reindex_partial_aux b fml mini is (i::iacc) vacc') +End + +Definition reindex_partial_def: + reindex_partial b fml mini is = + case mini of NONE => ([],[],is) + | SOME mini => + reindex_partial_aux b fml mini is [] [] +End + Definition subst_opt_subst_fun_def: subst_opt_subst_fun s c = subst_opt (subst_fun s) c End @@ -661,7 +682,7 @@ End (* Fast path for a special case *) Definition red_fast_def: red_fast s idopt pfs = ( - if s = Vector [] then + if s = INR (Vector []) then case idopt of NONE => NONE | SOME id =>( case pfs of @@ -686,11 +707,20 @@ Definition check_red_list_fast_def: NONE End +Definition get_earliest_def: + (get_earliest earliest (INR _) = SOME (0:num)) ∧ + (get_earliest earliest (INL (n,_)) = + sptree$lookup n earliest) +End + Definition check_red_list_def: - check_red_list ord obj b tcb fml inds id c s pfs idopt = + check_red_list ord obj b tcb fml inds id c s pfs idopt earliest = + let s = mk_subst s in case red_fast s idopt pfs of NONE => ( - let (rinds,fmlls) = reindex (b ∨ tcb) fml inds in + let mini = get_earliest earliest s in + let (rinds,fmlls,rest) = + reindex_partial (b ∨ tcb) fml mini inds in let nc = not c in let fml_not_c = update_resize fml NONE (SOME (nc,b)) id in let rsubs = red_subgoals ord (subst_fun s) c obj in @@ -704,34 +734,55 @@ Definition check_red_list_def: let rfml = rollback fml' id id' in if do_red_check idopt b tcb fml' s rfml rinds fmlls nc pfs rsubs then - SOME (rfml,rinds,id') + SOME (rfml,rinds ++ rest,id') else NONE)) | SOME (pf,cid) => check_red_list_fast b fml inds id c pf cid End +Definition min_opt_def: + min_opt i j = + case i of NONE => j + | SOME ii => MIN ii j +End + +(* v is the new index of the constraint (last arg) *) +Definition update_earliest_def: + (update_earliest earliest v [] = earliest) ∧ + (update_earliest earliest v ((i,n)::ns) = + update_earliest + (insert n (min_opt (lookup n earliest) v) earliest) + v + ns) +End + Definition opt_update_inds_def[simp]: - (opt_update_inds fml NONE id inds = (fml,inds,id)) ∧ - (opt_update_inds fml (SOME cc) id inds = + (opt_update_inds fml NONE id inds earliest = + (fml,inds,earliest,id)) ∧ + (opt_update_inds fml (SOME cc) id inds earliest = (update_resize fml NONE (SOME cc) id, sorted_insert id inds, + update_earliest earliest id (FST (FST cc)), id+1)) End Definition check_sstep_list_def: (check_sstep_list (sstep:sstep) ord obj tcb - (fml: (npbc # bool) option list) (inds:num list) (id:num) = + (fml: (npbc # bool) option list) (inds:num list) (id:num) + earliest = case sstep of | Lstep lstep => (case check_lstep_list lstep F fml 0 id of NONE => NONE | SOME (rfml,c,id') => - SOME(opt_update_inds rfml c id' inds)) + SOME (opt_update_inds rfml c id' inds earliest)) | Red c s pfs idopt => - case check_red_list ord obj F tcb fml inds id c s pfs idopt of + case check_red_list ord obj F tcb fml inds id c s pfs + idopt earliest of SOME (rfml,rinds,id') => - SOME( + SOME ( update_resize rfml NONE (SOME (c,tcb)) id', sorted_insert id' rinds, + update_earliest earliest id' (FST c), id'+1) | NONE => NONE) End @@ -1155,16 +1206,44 @@ Proof simp[lookup_core_only_list_def,any_el_ALT,EL_LUPDATE] QED +Definition earliest_rel_def: + earliest_rel fmlls earliest ⇔ + ∀x pos. + pos < min_opt (sptree$lookup x earliest) (LENGTH fmlls) ⇒ + case EL pos fmlls of NONE => T + | SOME c => ¬MEM x (MAP SND (FST (FST c))) +End + +(* If we already proved fml_rel, then we get earliest_rel + for free *) +Theorem fml_rel_fml_rel_earliest_rel: + fml_rel fml fmlls ∧ + earliest_rel fmlls earliest ∧ + fml_rel fml fmlls' ⇒ + earliest_rel fmlls' earliest +Proof + rw[fml_rel_def,earliest_rel_def]>> + first_x_assum(qspec_then `pos` mp_tac)>> + first_x_assum(qspecl_then [`x`,`pos`] mp_tac)>> + first_x_assum(qspec_then `pos` mp_tac)>> + rw[any_el_ALT]>> + gvs[min_opt_def]>> + every_case_tac>>fs[] +QED + Theorem fml_rel_check_red_list: fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ + earliest_rel fmlls earliest ∧ (∀n. n ≥ id ⇒ any_el n fmlls NONE = NONE) ∧ - check_red_list ord obj b tcb fmlls inds id c s pfs idopt = + check_red_list ord obj b tcb fmlls inds id c s pfs + idopt earliest = SOME (fmlls', inds', id') ⇒ check_red ord obj b tcb fml id c s pfs idopt = SOME id' ∧ fml_rel fml fmlls' ∧ ind_rel fmlls' inds' ∧ - (∀n. n ≥ id' ⇒ any_el n fmlls' NONE = NONE) ∧ + earliest_rel fmlls' earliest ∧ + (∀n. n ≥ id ⇒ any_el n fmlls' NONE = NONE) ∧ id ≤ id' Proof strip_tac>> @@ -1195,6 +1274,8 @@ Proof gvs[do_red_check_def,AllCaseEqs(),insert_fml_def]>> TOP_CASE_TAC>>fs[] >- ( + cheat + (* rpt (pairarg_tac>>fs[])>> (drule_at Any) split_goals_hash_imp_split_goals>> disch_then (qspec_then`mk_core_fml (b ∨ tcb) fml` mp_tac)>> @@ -1253,16 +1334,26 @@ Proof drule (GSYM fml_rel_lookup_core_only)>> strip_tac>>fs[]>> gvs[lookup_core_only_list_def,AllCaseEqs()])>> - fs[])>> + fs[]*))>> match_mp_tac (GEN_ALL fml_rel_check_contradiction_fml)>> metis_tac[])>> - CONJ_TAC>- ( + CONJ_ASM1_TAC>- ( match_mp_tac fml_rel_rollback>>rw[]>>fs[])>> CONJ_TAC >- ( + cheat + (* match_mp_tac ind_rel_rollback_2>> simp[]>> - metis_tac[ind_rel_reindex])>> - simp[rollback_def,any_el_list_delete_list,MEM_MAP,MEM_COUNT_LIST])>> + metis_tac[ind_rel_reindex]*))>> + CONJ_TAC >- + metis_tac[fml_rel_fml_rel_earliest_rel]>> + simp[rollback_def,any_el_list_delete_list,MEM_MAP,MEM_COUNT_LIST]>> + rw[]>> + qsuff_tac `n < id'` + >- ( + simp[]>> + intLib.ARITH_TAC)>> + CCONTR_TAC>>gvs[])>> gvs[check_red_list_fast_def,AllCaseEqs(),check_red_def,red_fast_def,extract_clauses_def,check_subproofs_def,insert_fml_def,check_lstep_list_def] >- ( drule fml_rel_update_resize>> @@ -1270,15 +1361,20 @@ Proof drule_all fml_rel_check_contradiction_fml>> simp[]>> strip_tac>> - rw[] + CONJ_ASM1_TAC >- ( match_mp_tac fml_rel_rollback>> simp[]>> - rw[any_el_update_resize]) + rw[any_el_update_resize])>> + CONJ_TAC >- ( match_mp_tac ind_rel_rollback_2>> simp[any_el_update_resize])>> - simp[rollback_def,any_el_list_delete_list,MEM_MAP,MEM_COUNT_LIST,any_el_update_resize])>> + CONJ_TAC + >- (* earliest_rel *) + metis_tac[fml_rel_fml_rel_earliest_rel] + >- + simp[rollback_def,any_el_list_delete_list,MEM_MAP,MEM_COUNT_LIST,any_el_update_resize])>> `fml_rel (insert id ((not c,b)) fml) (update_resize fmlls NONE (SOME (not c,b)) id)` by metis_tac[fml_rel_update_resize]>> @@ -1292,7 +1388,7 @@ Proof simp[any_el_update_resize]>> ntac 3 strip_tac>> CONJ_TAC >- metis_tac[fml_rel_check_contradiction_fml]>> - CONJ_TAC >- ( + CONJ_ASM1_TAC >- ( match_mp_tac fml_rel_rollback>> fs[]>> rw[]>- metis_tac[]>> @@ -1300,11 +1396,20 @@ Proof CONJ_TAC >- ( match_mp_tac ind_rel_rollback_2>> simp[])>> - simp[rollback_def,any_el_list_delete_list,MEM_MAP,MEM_COUNT_LIST,any_el_update_resize] + CONJ_TAC >- (* earliest_rel *) + metis_tac[fml_rel_fml_rel_earliest_rel]>> + simp[rollback_def,any_el_list_delete_list,MEM_MAP,MEM_COUNT_LIST,any_el_update_resize]>> + rw[]>> + qsuff_tac `n < id'` + >- ( + simp[]>> + intLib.ARITH_TAC)>> + CCONTR_TAC>>gvs[] QED Theorem opt_update_inds_opt_update: - opt_update_inds fml c id inds = (fml',inds',id') ⇒ + opt_update_inds fml c id inds earliest = + (fml',inds',earliest',id') ⇒ opt_update fml c id = (fml',id') Proof Cases_on`c`>>rw[] @@ -1329,24 +1434,75 @@ QED Theorem opt_update_inds_ind_rel: ind_rel fml inds ∧ - opt_update_inds fml c id inds = (fml',inds',id') ⇒ + opt_update_inds fml c id inds earliest = + (fml',inds',earliest',id') ⇒ ind_rel fml' inds' Proof Cases_on`c`>>rw[]>>fs[]>> simp[ind_rel_update_resize_sorted_insert] QED +Theorem earliest_rel_check_lstep_list: + earliest_rel fmlls earliest ∧ + (∀n. n ≥ id ⇒ any_el n fmlls NONE = NONE) ∧ + check_lstep_list lstep b fmlls mindel id = + SOME (fmlls',x,y) ⇒ + earliest_rel fmlls' earliest +Proof + rw[]>> + fs[earliest_rel_def]>> + rw[]>> + first_x_assum(qspecl_then[`x'`,`pos`] mp_tac)>> + drule (CONJUNCT1 check_lstep_list_id_del)>> + drule (CONJUNCT1 check_lstep_list_id_upper)>> + disch_then drule>> + rw[]>> + Cases_on`pos < id`>>fs[] + >- ( + gvs[min_opt_def]>> + every_case_tac>>gvs[]>> + first_x_assum drule >> + rw[any_el_ALT]>> + gvs[])>> + `pos ≥ id` by fs[]>> + gvs[min_opt_def]>> + every_case_tac>>gvs[]>> + first_x_assum drule >> + rw[any_el_ALT]>> + gvs[] +QED + +Theorem earliest_rel_update_resize_update_earliest: + earliest_rel fml earliest ⇒ + earliest_rel (update_resize fmlls NONE (SOME v) n) + (update_earliest earliest n (FST (FST v))) +Proof + cheat +QED + +Theorem opt_update_inds_earliest_rel: + earliest_rel fml earliest ∧ + opt_update_inds fml c id inds earliest = + (fml',inds',earliest',id') ⇒ + earliest_rel fml' earliest' +Proof + Cases_on`c`>>rw[]>>fs[]>> + metis_tac[earliest_rel_update_resize_update_earliest] +QED + Theorem fml_rel_check_sstep_list: ∀sstep ord obj fmlls inds id fmlls' id' inds' fml. fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ + earliest_rel fmlls earliest ∧ (∀n. n ≥ id ⇒ any_el n fmlls NONE = NONE) ∧ - check_sstep_list sstep ord obj tcb fmlls inds id = - SOME (fmlls',inds',id') ⇒ + check_sstep_list sstep ord obj tcb fmlls inds id earliest = + SOME (fmlls',inds',earliest',id') ⇒ ∃fml'. check_sstep sstep ord obj tcb fml id = SOME(fml',id') ∧ fml_rel fml' fmlls' ∧ ind_rel fmlls' inds' ∧ + earliest_rel fmlls' earliest' ∧ (∀n. n ≥ id' ⇒ any_el n fmlls' NONE = NONE) ∧ id ≤ id' Proof @@ -1360,16 +1516,14 @@ Proof rw[]>>simp[]>> CONJ_TAC >- metis_tac[opt_update_inds_ind_rel,ind_rel_check_lstep_list]>> - CONJ_TAC - >- ( - drule (CONJUNCT1 check_lstep_list_id_upper)>> - drule opt_update_id_upper>> - drule opt_update_id>> - drule (CONJUNCT1 check_lstep_list_id)>> - simp[])>> - drule (CONJUNCT1 check_lstep_list_id)>> + CONJ_TAC >- + metis_tac[opt_update_inds_earliest_rel,earliest_rel_check_lstep_list]>> + drule (CONJUNCT1 check_lstep_list_id_upper)>> + drule opt_update_id_upper>> drule opt_update_id>> - simp[]) + drule (CONJUNCT1 check_lstep_list_id)>> + drule (CONJUNCT1 check_lstep_list_mindel)>> + simp[]>>rw[]) >- ( (* Red*) gvs[AllCaseEqs(),insert_fml_def]>> drule_all fml_rel_check_red_list>> @@ -1377,7 +1531,9 @@ Proof >- metis_tac[fml_rel_update_resize] >- - metis_tac[ind_rel_update_resize_sorted_insert]>> + metis_tac[ind_rel_update_resize_sorted_insert] + >- + metis_tac[PAIR,FST,SND,earliest_rel_update_resize_update_earliest]>> simp[any_el_update_resize]) QED @@ -1476,7 +1632,7 @@ Definition do_change_obj_check_def: End Definition emp_vec_def: - emp_vec = Vector [] + emp_vec = INR (Vector []) End Definition check_change_obj_list_def: @@ -1497,6 +1653,7 @@ Definition check_change_obj_list_def: else NONE) End +(* TODO *) Definition check_cstep_list_def: check_cstep_list cstep fml inds pc = case cstep of @@ -1509,6 +1666,7 @@ Definition check_cstep_list_def: let (rinds,fmlls) = reindex F fml inds in let corels = core_fmlls fml rinds in let fml_not_c = update_resize fml NONE (SOME (nc,F)) id in + let s = mk_subst s in let w = subst_fun s in let dsubs = dom_subgoals spo w c pc.obj in case extract_clauses_list s F fml dsubs pfs [] of diff --git a/examples/pseudo_bool/npbc_checkScript.sml b/examples/pseudo_bool/npbc_checkScript.sml index f90a8f2615..2a5d9e2551 100644 --- a/examples/pseudo_bool/npbc_checkScript.sml +++ b/examples/pseudo_bool/npbc_checkScript.sml @@ -1,7 +1,7 @@ (* Pseudo-boolean constraints proof format and checker *) -open preamble npbcTheory mlstringTheory mlintTheory mlvectorTheory; +open preamble npbcTheory mlstringTheory mlintTheory mlvectorTheory spt_to_vecTheory; val _ = new_theory "npbc_check"; @@ -583,13 +583,13 @@ Proof QED *) -Type subst = ``:(bool + num lit) option vector``; +Type subst_raw = ``:(num , bool + num lit) alist``; (* Steps that preserve satisfiability *) Datatype: sstep = | Lstep lstep (* Id representing a contradiction *) - | Red npbc subst (( ((num + num) # num) option, (lstep list)) alist) (num option) + | Red npbc subst_raw (( ((num + num) # num) option, (lstep list)) alist) (num option) (* the alist represents a subproof NONE -> top level step SOME (INL n,id) -> database proofgoals, contradiction at id @@ -668,8 +668,13 @@ Definition check_subproofs_def: | res => NONE)) End +Type subst = ``:(num # (bool + num lit)) + (bool + num lit) option vector``; + Definition subst_fun_def: subst_fun (s:subst) n = + case s of + INL (m,v) => if n = m then SOME v else NONE + | INR s => if n < length s then sub s n else NONE @@ -766,12 +771,18 @@ Definition mk_core_fml_def: if b ⇒ b' then SOME x else NONE) fml End +Definition mk_subst_def: + (mk_subst [(n,v)] = INL (n,v)) ∧ + (mk_subst xs = INR (spt_to_vec (fromAList xs))) +End + (* The tcb flag indicates we're in to-core mode where it is guaranteed that the core formula implies derived *) Definition check_red_def: check_red ord obj b tcb fml id c s pfs idopt = ( let nc = not c in let (fml_not_c,id1) = insert_fml b fml id (not c) in + let s = mk_subst s in let w = subst_fun s in let rsubs = red_subgoals ord w c obj in case extract_clauses w b fml rsubs pfs [] of @@ -1203,7 +1214,7 @@ Proof \\ pairarg_tac \\ fs[] \\ match_mp_tac (GEN_ALL substitution_redundancy_obj_po) \\ simp[] - \\ qexists_tac ‘subst_fun s’ \\ fs [] + \\ qexists_tac ‘subst_fun (mk_subst s)’ \\ fs [] \\ fs[EVERY_MEM,MEM_MAP,EXISTS_PROD] \\ `id ∉ domain fml` by fs[id_ok_def] \\ @@ -1234,7 +1245,7 @@ Proof >- ( (* objective *) Cases_on`obj`>> - last_x_assum(qspec_then`SUC(LENGTH (dom_subst (subst_fun s) ord))` assume_tac)>> + last_x_assum(qspec_then`SUC(LENGTH (dom_subst (subst_fun (mk_subst s)) ord))` assume_tac)>> fs[]>> drule_all lookup_extract_pids_r>>rw[] \\ drule extract_clauses_MEM_INR @@ -1264,7 +1275,7 @@ Proof \\ metis_tac[INSERT_SING_UNION,UNION_COMM]) (* rest of redundancy *) \\ gvs [GSYM unsat_iff_implies] - \\ Cases_on ‘subst_opt (subst_fun s) x'’ \\ fs [] + \\ Cases_on ‘subst_opt (subst_fun (mk_subst s)) x'’ \\ fs [] >- ( imp_res_tac subst_opt_NONE \\ CCONTR_TAC \\ gvs [satisfiable_def,not_thm] @@ -1276,7 +1287,7 @@ Proof \\ strip_tac \\ rename1`lookup nn _ = SOME xx` \\ rename1`subst_opt _ _ = SOME yy` - \\ `MEM (nn,yy) (toAList (map_opt (subst_opt (subst_fun s)) + \\ `MEM (nn,yy) (toAList (map_opt (subst_opt (subst_fun (mk_subst s))) (mk_core_fml (b ∨ tcb) fml)))` by simp[MEM_toAList,lookup_map_opt] \\ drule_all split_goals_checked \\ rw[] @@ -1523,12 +1534,12 @@ QED Datatype: cstep = (* Derivation steps *) - | Dom npbc subst (( ((num + num) # num) option, (lstep list)) alist) (num option) + | Dom npbc subst_raw (( ((num + num) # num) option, (lstep list)) alist) (num option) | Sstep sstep (* Deletion steps *) | CheckedDelete num - subst + subst_raw (( ((num + num) # num) option, (lstep list)) alist) (num option) | UncheckedDelete (num list) @@ -1987,6 +1998,7 @@ Definition check_cstep_def: | SOME spo => ( let nc = not c in let (fml_not_c,id1) = insert_fml F fml pc.id (not c) in + let s = mk_subst s in let w = subst_fun s in let dsubs = dom_subgoals spo w c pc.obj in case extract_clauses w F fml dsubs pfs [] of @@ -2380,7 +2392,7 @@ Proof PairCases_on`x`>> match_mp_tac (GEN_ALL good_spo_dominance)>> simp[]>> - qexists_tac ‘subst_fun v’>>fs[]>> + qexists_tac ‘subst_fun (mk_subst l)’>>fs[]>> CONJ_TAC >- metis_tac[core_only_fml_T_SUBSET_F]>> CONJ_TAC >- @@ -2393,7 +2405,7 @@ Proof (* core constraints*) rw[sat_implies_def,satisfies_def] \\ gs[PULL_EXISTS,GSYM range_mk_core_fml,range_def,lookup_mk_core_fml] - \\ Cases_on ‘subst_opt (subst_fun v) x’ \\ fs [] + \\ Cases_on ‘subst_opt (subst_fun (mk_subst l)) x’ \\ fs [] THEN1 ( imp_res_tac subst_opt_NONE \\ gvs[lookup_core_only_def,AllCaseEqs()] @@ -2401,7 +2413,7 @@ Proof \\ fs [satisfies_def,range_def,PULL_EXISTS] \\ metis_tac[]) \\ rename1`subst_opt _ _ = SOME yy` - \\ `MEM (n,yy) (toAList (map_opt (subst_opt (subst_fun v)) + \\ `MEM (n,yy) (toAList (map_opt (subst_opt (subst_fun (mk_subst l))) (mk_core_fml T fml)))` by simp[MEM_toAList,lookup_map_opt,lookup_mk_core_fml,lookup_core_only_def] \\ drule_all split_goals_checked \\ rw[] @@ -2432,7 +2444,8 @@ Proof gvs[lookup_mk_core_fml]>> drule lookup_core_only_T_imp_F>> rw[]) - \\ `subst (subst_fun v) c = subst (subst_fun v) x` by + \\ `subst (subst_fun (mk_subst l)) c = + subst (subst_fun (mk_subst l)) x` by metis_tac[subst_opt_SOME] \\ metis_tac[])>> CONJ_TAC >- ( @@ -2459,7 +2472,7 @@ Proof CONJ_TAC >- ( (* negated order constraint *) fs[core_only_fml_def]>> - last_x_assum(qspec_then`LENGTH (dom_subst (subst_fun v) (SOME ((x0,x1,x2),x3)))` assume_tac)>> + last_x_assum(qspec_then`LENGTH (dom_subst (subst_fun (mk_subst l)) (SOME ((x0,x1,x2),x3)))` assume_tac)>> gs[ADD1]>> drule_all lookup_extract_pids_r>> simp[]>> rw[] @@ -2476,7 +2489,7 @@ Proof fs[core_only_fml_def]>> Cases_on`pc.obj`>> simp[]>> - last_x_assum(qspec_then`SUC(LENGTH (dom_subst (subst_fun v) (SOME ((x0,x1,x2),x3))))` assume_tac)>> + last_x_assum(qspec_then`SUC(LENGTH (dom_subst (subst_fun (mk_subst l)) (SOME ((x0,x1,x2),x3))))` assume_tac)>> gs[ADD1]>> drule_all lookup_extract_pids_r>> simp[]>>rw[] diff --git a/examples/pseudo_bool/pb_parseScript.sml b/examples/pseudo_bool/pb_parseScript.sml index fe0957833d..84fe2f13aa 100644 --- a/examples/pseudo_bool/pb_parseScript.sml +++ b/examples/pseudo_bool/pb_parseScript.sml @@ -1,7 +1,7 @@ (* Parse and print for pbc, npb_check *) -open preamble pbcTheory pbc_normaliseTheory npbc_checkTheory spt_to_vecTheory; +open preamble pbcTheory pbc_normaliseTheory npbc_checkTheory; val _ = new_theory "pb_parse"; @@ -295,24 +295,24 @@ Definition parse_subst_aux_def: (parse_subst_aux f_ns (INL v :: INL arr :: r ::rest) acc = if arr = strlit "->" then case parse_var f_ns v of - NONE => ([],LN,f_ns) + NONE => ([],[],f_ns) | SOME (v,f_ns) => (case r of INR r => - if r = 0:int then parse_subst_aux f_ns rest (insert v (INL F) acc) - else if r = 1 then parse_subst_aux f_ns rest (insert v (INL T) acc) - else ([],LN,f_ns) + if r = 0:int then parse_subst_aux f_ns rest ((v,INL F)::acc) + else if r = 1 then parse_subst_aux f_ns rest ((v,INL T)::acc) + else ([],[],f_ns) | INL r => - (case parse_lit_num f_ns r of NONE => ([],LN,f_ns) - | SOME (l, f_ns) => parse_subst_aux f_ns rest (insert v (INR l) acc))) - else ([],LN,f_ns)) ∧ + (case parse_lit_num f_ns r of NONE => ([],[],f_ns) + | SOME (l, f_ns) => parse_subst_aux f_ns rest ((v,INR l)::acc))) + else ([],[],f_ns)) ∧ (parse_subst_aux f_ns ls acc = (ls, acc,f_ns)) End Definition parse_subst_def: parse_subst f_ns ls = - let (ls, acc,f_ns) = parse_subst_aux f_ns ls LN in - (ls, spt_to_vec acc, f_ns) + let (ls, acc,f_ns) = parse_subst_aux f_ns ls [] in + (ls, acc, f_ns) End Definition map_f_ns_def: @@ -408,8 +408,8 @@ Definition check_end_def: | _ => NONE End -Definition is_empty_vec_def: - is_empty_vec v ⇔ v = Vector [] +Definition is_empty_def: + is_empty v ⇔ v = [] End (* Parsing lsteps in a subproof until parsing @@ -427,7 +427,7 @@ Definition parse_lsteps_aux_def: | SOME (INL step, f_ns') => parse_lsteps_aux f_ns' ss (step::acc) | SOME (INR (c,s), f_ns') => - if ¬(is_empty_vec s) then NONE + if ¬(is_empty s) then NONE else case parse_lsteps_aux f_ns' ss [] of NONE => NONE @@ -466,7 +466,7 @@ Theorem parse_lsteps_aux_thm: | SOME (INL step, f_ns') => parse_lsteps_aux f_ns' ss (step::acc) | SOME (INR (c,s), f_ns') => - if ¬(is_empty_vec s) then NONE + if ¬(is_empty s) then NONE else case parse_lsteps_aux f_ns' ss [] of NONE => NONE @@ -653,7 +653,7 @@ End (* Check special cases whether a Red step can be converted to a Con step *) Definition reduce_pf_def: reduce_pf s pf res = - if is_empty_vec s then + if is_empty s then case res of NONE => NONE | SOME id =>( case pf of @@ -876,9 +876,9 @@ strlit"end"; Datatype: par = | Done cstep - | Dompar npbc subst + | Dompar npbc subst_raw | StoreOrderpar mlstring - | CheckedDeletepar num subst + | CheckedDeletepar num subst_raw | ChangeObjpar bool ((int # num) list # int) End From 49969be4edc2d6ce28d5013e71a8a0c2e13e8a85 Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sun, 21 Jan 2024 22:34:41 +0800 Subject: [PATCH 03/38] reduce cheats --- .../pseudo_bool/array/npbc_listScript.sml | 371 ++++++++++++++---- 1 file changed, 297 insertions(+), 74 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 243a7355ff..958da406ac 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -79,10 +79,6 @@ Definition sorted_insert_def: else y::(sorted_insert x ys)) End -Definition full_sorted_insert_def: - full_sorted_insert n = OPTION_MAP (sorted_insert n) -End - Definition check_contradiction_fml_list_def: check_contradiction_fml_list b fml n = case lookup_core_only_list b fml n of @@ -1101,6 +1097,17 @@ Proof metis_tac[] QED +Theorem SORTED_reindex: + SORTED $>= inds ∧ + reindex b fml inds = (is,vs) ⇒ + SORTED $>= is +Proof + rw[]>>drule FST_reindex_characterize>> + rw[]>> + match_mp_tac SORTED_FILTER>> + fs[transitive_def] +QED + Theorem ind_rel_reindex: ind_rel fml inds ∧ reindex b fml inds = (is,vs) ⇒ @@ -1110,6 +1117,101 @@ Proof fs[ind_rel_def,MEM_FILTER] QED +Theorem SORTED_HEAD_LESS: + ¬(h ≥ mini:num) ∧ + SORTED $>= (h::inds) ⇒ + EVERY (λx. x < mini) inds +Proof + DEP_REWRITE_TAC [SORTED_EQ]>> + simp[transitive_def,EVERY_MEM]>> + rw[]>> + first_x_assum drule>> + fs[] +QED + +Theorem reindex_partial_aux: + ∀inds iacc vacc. + SORTED $>= inds ⇒ + reindex_partial_aux b fmlls mini inds iacc vacc = + let finds = FILTER (λx. x ≥ mini) inds in + let binds = FILTER (λx. x < mini) inds in + let is = FILTER (λx. IS_SOME (any_el x fmlls NONE)) finds in + let vs = + MAP (λx. THE (lookup_core_only_list b fmlls x)) + (FILTER (λx. IS_SOME (lookup_core_only_list b fmlls x)) + finds) in + (REVERSE iacc ++ is, REVERSE vs ++ vacc, binds) +Proof + Induct>>simp[reindex_partial_aux_def]>> + rw[]>>fs[] + >- ( + drule_all SORTED_HEAD_LESS>>rw[]>> + fs[FILTER_FILTER,FILTER_EQ_NIL,EVERY_MEM,MEM_FILTER]>> + rw[]>> + first_x_assum drule>>fs[]) + >- ( + drule_all SORTED_HEAD_LESS>>rw[]>> + fs[FILTER_FILTER,FILTER_EQ_NIL,EVERY_MEM,MEM_FILTER]>> + rw[]>> + first_x_assum drule>>fs[]) + >- ( + drule_all SORTED_HEAD_LESS>>rw[]>> + metis_tac[GSYM FILTER_EQ_ID])>> + drule SORTED_TL>>strip_tac>>fs[]>> + every_case_tac>> + gvs[lookup_core_only_list_def,IS_SOME_EXISTS,AllCaseEqs()] +QED + +Theorem FST_reindex_partial_characterize: + SORTED $>= inds ∧ + reindex_partial b fmlls mini inds = (is,vs,rest) ⇒ + case mini of NONE => is = [] + | SOME mini => + is = FILTER (λx. IS_SOME (any_el x fmlls NONE)) + (FILTER (λx. x ≥ mini) inds) +Proof + rw[reindex_partial_def,reindex_partial_aux]>> + every_case_tac>>fs[]>> + drule reindex_partial_aux>> + rw[]>>gvs[] +QED + +Theorem SND_reindex_partial_characterize: + fml_rel fml fmlls ∧ + SORTED $>= inds ∧ + reindex_partial b fmlls mini inds = (is,vs,rest) ⇒ + set vs ⊆ core_only_fml b fml +Proof + rw[reindex_partial_def]>> + gvs[AllCaseEqs()]>> + drule reindex_partial_aux>> + rw[]>>gvs[]>> + simp[SUBSET_DEF,MEM_MAP,MEM_FILTER,PULL_EXISTS,range_def]>> + rw[]>> + fs[IS_SOME_EXISTS]>> + drule fml_rel_lookup_core_only>> + rw[]>> + gvs[lookup_core_only_def,core_only_fml_def,AllCaseEqs()] + >- + metis_tac[]>> + rw[]>>fs[]>> + metis_tac[] +QED + +Theorem ind_rel_reindex_partial: + SORTED $>= inds ∧ + ind_rel fml inds ∧ + reindex_partial b fml mini inds = (is,vs,rest) ⇒ + ind_rel fml (is++rest) ∧ + SORTED $>= (is++rest) +Proof + strip_tac>> + gvs[reindex_partial_def,AllCaseEqs()]>> + drule reindex_partial_aux>> + strip_tac>>gvs[]>> + cheat +QED + Theorem MEM_subst_indexes: ∀inds i c. MEM i inds ∧ @@ -1234,6 +1336,7 @@ QED Theorem fml_rel_check_red_list: fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ + SORTED $>= inds ∧ earliest_rel fmlls earliest ∧ (∀n. n ≥ id ⇒ any_el n fmlls NONE = NONE) ∧ check_red_list ord obj b tcb fmlls inds id c s pfs @@ -1242,8 +1345,9 @@ Theorem fml_rel_check_red_list: check_red ord obj b tcb fml id c s pfs idopt = SOME id' ∧ fml_rel fml fmlls' ∧ ind_rel fmlls' inds' ∧ + SORTED $>= inds' ∧ earliest_rel fmlls' earliest ∧ - (∀n. n ≥ id ⇒ any_el n fmlls' NONE = NONE) ∧ + (∀n. n ≥ id' ⇒ any_el n fmlls' NONE = NONE) ∧ id ≤ id' Proof strip_tac>> @@ -1274,21 +1378,23 @@ Proof gvs[do_red_check_def,AllCaseEqs(),insert_fml_def]>> TOP_CASE_TAC>>fs[] >- ( - cheat - (* rpt (pairarg_tac>>fs[])>> (drule_at Any) split_goals_hash_imp_split_goals>> disch_then (qspec_then`mk_core_fml (b ∨ tcb) fml` mp_tac)>> impl_tac >- ( simp[range_mk_core_fml]>> - match_mp_tac (GEN_ALL SND_reindex_characterize)>> + match_mp_tac (GEN_ALL SND_reindex_partial_characterize)>> metis_tac[])>> match_mp_tac split_goals_same_goals>> simp[EXTENSION,FORALL_PROD]>> rw[]>>eq_tac>>rw[] >- ( fs[MEM_toAList,lookup_map_opt,AllCaseEqs()]>> - match_mp_tac (GEN_ALL MEM_subst_indexes)>> + (* Prove that every goal is covered, + use the fact that subst_opt ... = SOME + and earliest_rel *) + cheat + (*match_mp_tac (GEN_ALL MEM_subst_indexes)>> gvs[lookup_mk_core_fml]>> first_x_assum (irule_at Any)>> `∃b'. @@ -1317,11 +1423,11 @@ Proof drule (GSYM fml_rel_lookup_core_only)>> strip_tac>>fs[]>> gvs[lookup_core_only_list_def,AllCaseEqs()]>> - metis_tac[]))>> + metis_tac[])*))>> drule subst_indexes_MEM>> rw[MEM_toAList,lookup_map_opt]>> - drule FST_reindex_characterize>> - strip_tac>>gvs[]>> + drule_all FST_reindex_partial_characterize>> + TOP_CASE_TAC>>rw[]>>gvs[]>> fs[rollback_def,lookup_core_only_list_list_delete_list,MEM_MAP,MEM_COUNT_LIST,MEM_FILTER]>> `p_1 < id` by ( CCONTR_TAC>>fs[]>> @@ -1334,26 +1440,21 @@ Proof drule (GSYM fml_rel_lookup_core_only)>> strip_tac>>fs[]>> gvs[lookup_core_only_list_def,AllCaseEqs()])>> - fs[]*))>> + fs[])>> match_mp_tac (GEN_ALL fml_rel_check_contradiction_fml)>> metis_tac[])>> CONJ_ASM1_TAC>- ( match_mp_tac fml_rel_rollback>>rw[]>>fs[])>> CONJ_TAC >- ( - cheat - (* match_mp_tac ind_rel_rollback_2>> - simp[]>> - metis_tac[ind_rel_reindex]*))>> + simp[] >> + metis_tac[ind_rel_reindex_partial])>> + CONJ_TAC >- + metis_tac[ind_rel_reindex_partial]>> CONJ_TAC >- metis_tac[fml_rel_fml_rel_earliest_rel]>> simp[rollback_def,any_el_list_delete_list,MEM_MAP,MEM_COUNT_LIST]>> - rw[]>> - qsuff_tac `n < id'` - >- ( - simp[]>> - intLib.ARITH_TAC)>> - CCONTR_TAC>>gvs[])>> + rw[])>> gvs[check_red_list_fast_def,AllCaseEqs(),check_red_def,red_fast_def,extract_clauses_def,check_subproofs_def,insert_fml_def,check_lstep_list_def] >- ( drule fml_rel_update_resize>> @@ -1399,12 +1500,7 @@ Proof CONJ_TAC >- (* earliest_rel *) metis_tac[fml_rel_fml_rel_earliest_rel]>> simp[rollback_def,any_el_list_delete_list,MEM_MAP,MEM_COUNT_LIST,any_el_update_resize]>> - rw[]>> - qsuff_tac `n < id'` - >- ( - simp[]>> - intLib.ARITH_TAC)>> - CCONTR_TAC>>gvs[] + rw[] QED Theorem opt_update_inds_opt_update: @@ -1432,6 +1528,16 @@ Proof metis_tac[] QED +Theorem opt_update_inds_SORTED: + SORTED $>= inds ∧ + opt_update_inds fml c id inds earliest = + (fml',inds',earliest',id') ⇒ + SORTED $>= inds' +Proof + Cases_on`c`>>rw[]>>fs[]>> + metis_tac[SORTED_sorted_insert] +QED + Theorem opt_update_inds_ind_rel: ind_rel fml inds ∧ opt_update_inds fml c id inds earliest = @@ -1472,10 +1578,11 @@ Proof gvs[] QED +(* TODO: maybe some induction? *) Theorem earliest_rel_update_resize_update_earliest: earliest_rel fml earliest ⇒ - earliest_rel (update_resize fmlls NONE (SOME v) n) - (update_earliest earliest n (FST (FST v))) + earliest_rel (update_resize fml NONE (SOME (v,b)) n) + (update_earliest earliest n (FST v)) Proof cheat QED @@ -1487,13 +1594,14 @@ Theorem opt_update_inds_earliest_rel: earliest_rel fml' earliest' Proof Cases_on`c`>>rw[]>>fs[]>> - metis_tac[earliest_rel_update_resize_update_earliest] + metis_tac[earliest_rel_update_resize_update_earliest,FST,PAIR] QED Theorem fml_rel_check_sstep_list: ∀sstep ord obj fmlls inds id fmlls' id' inds' fml. fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ + SORTED $>= inds ∧ earliest_rel fmlls earliest ∧ (∀n. n ≥ id ⇒ any_el n fmlls NONE = NONE) ∧ check_sstep_list sstep ord obj tcb fmlls inds id earliest = @@ -1502,6 +1610,7 @@ Theorem fml_rel_check_sstep_list: check_sstep sstep ord obj tcb fml id = SOME(fml',id') ∧ fml_rel fml' fmlls' ∧ ind_rel fmlls' inds' ∧ + SORTED $>= inds' ∧ earliest_rel fmlls' earliest' ∧ (∀n. n ≥ id' ⇒ any_el n fmlls' NONE = NONE) ∧ id ≤ id' @@ -1516,6 +1625,8 @@ Proof rw[]>>simp[]>> CONJ_TAC >- metis_tac[opt_update_inds_ind_rel,ind_rel_check_lstep_list]>> + CONJ_TAC >- + metis_tac[opt_update_inds_SORTED]>> CONJ_TAC >- metis_tac[opt_update_inds_earliest_rel,earliest_rel_check_lstep_list]>> drule (CONJUNCT1 check_lstep_list_id_upper)>> @@ -1532,6 +1643,8 @@ Proof metis_tac[fml_rel_update_resize] >- metis_tac[ind_rel_update_resize_sorted_insert] + >- + metis_tac[SORTED_sorted_insert] >- metis_tac[PAIR,FST,SND,earliest_rel_update_resize_update_earliest]>> simp[any_el_update_resize]) @@ -1653,9 +1766,8 @@ Definition check_change_obj_list_def: else NONE) End -(* TODO *) Definition check_cstep_list_def: - check_cstep_list cstep fml inds pc = + check_cstep_list cstep fml inds earliest pc = case cstep of Dom c s pfs idopt => (case pc.ord of @@ -1680,13 +1792,14 @@ Definition check_cstep_list_def: SOME( update_resize rfml NONE (SOME (c,pc.tcb)) id', sorted_insert id' rinds, + update_earliest earliest id' (FST c), pc with id := id'+1) else NONE))) | Sstep sstep => (case check_sstep_list sstep pc.ord pc.obj pc.tcb - fml inds pc.id of - SOME(fml',inds',id') => - SOME(fml',inds', pc with id := id') + fml inds pc.id earliest of + SOME(fml',inds',earliest',id') => + SOME(fml',inds', earliest',pc with id := id') | NONE => NONE) | CheckedDelete n s pfs idopt => ( if check_tcb_idopt pc.tcb idopt then @@ -1695,33 +1808,35 @@ Definition check_cstep_list_def: | SOME c => (let nfml = delete_list n fml in case check_red_list pc.ord pc.obj T pc.tcb - nfml inds pc.id c s pfs idopt of + nfml inds pc.id c s pfs idopt earliest of SOME (ncf',inds',id') => - SOME (ncf', inds', pc with <| id := id' |>) + SOME (ncf', inds', earliest, pc with <| id := id' |>) | NONE => NONE) ) else NONE) | UncheckedDelete ls => ( (* Either no order or all ids are in core *) if ¬pc.tcb ∧ pc.ord = NONE then - SOME (list_delete_list ls fml, inds, pc with chk := F) + SOME (list_delete_list ls fml, inds, + earliest, pc with chk := F) else case all_core_list fml inds [] of NONE => NONE | SOME inds' => - SOME (list_delete_list ls fml, inds', pc with chk := F)) + SOME (list_delete_list ls fml, inds', + earliest, pc with chk := F)) | Transfer ls => (case core_from_inds fml ls of NONE => NONE | SOME fml' => - SOME (fml', inds, pc)) + SOME (fml', inds, earliest, pc)) | StrengthenToCore b => (let inds' = FST (reindex F fml inds) in let pc' = pc with tcb := b in if b then (case core_from_inds fml inds' of NONE => NONE - | SOME fml' => SOME (fml',inds',pc')) + | SOME fml' => SOME (fml',inds', earliest, pc')) else - SOME (fml,inds',pc')) + SOME (fml,inds',earliest,pc')) | LoadOrder nn xs => (let inds' = FST (reindex F fml inds) in case ALOOKUP pc.orders nn of NONE => NONE @@ -1729,19 +1844,20 @@ Definition check_cstep_list_def: if LENGTH xs = LENGTH (FST (SND ord')) then case core_from_inds fml inds' of NONE => NONE | SOME fml' => - SOME (fml',inds',pc with ord := SOME (ord',xs)) + SOME (fml',inds',earliest,pc with ord := SOME (ord',xs)) else NONE) | UnloadOrder => (case pc.ord of NONE => NONE | SOME spo => - SOME (fml,inds, pc with ord := NONE)) + SOME (fml,inds, earliest, pc with ord := NONE)) | StoreOrder nn spo ws pfsr pfst => if check_good_ord spo ∧ check_ws spo ws then case check_transitivity spo ws pfst of NONE => NONE | SOME id => if check_reflexivity spo pfsr id then - SOME (fml, inds, pc with orders := (nn,spo)::pc.orders) + SOME (fml, inds, + earliest, pc with orders := (nn,spo)::pc.orders) else NONE else NONE @@ -1757,12 +1873,13 @@ Definition check_cstep_list_def: SOME ( update_resize fml NONE (SOME (c,T)) pc.id, sorted_insert pc.id inds, + update_earliest earliest pc.id (FST c), pc with <| id := pc.id+1; bound := bound'; dbound := dbound' |>) else - SOME (fml, inds, + SOME (fml, inds, earliest, pc with <| bound := bound'; dbound := dbound' |>)) @@ -1771,11 +1888,11 @@ Definition check_cstep_list_def: NONE => NONE | SOME (fml',fc',id') => SOME ( - fml', inds, + fml', inds, earliest, pc with <| id:=id'; obj:=SOME fc' |>)) | CheckObj fc' => if check_eq_obj pc.obj fc' - then SOME (fml, inds, pc) + then SOME (fml, inds, earliest, pc) else NONE End @@ -1876,16 +1993,40 @@ Proof simp[rollback_def,any_el_list_delete_list,MEM_MAP,MEM_COUNT_LIST] QED +(* should be easy *) +Theorem earliest_rel_list_delete_list: + ∀l fmlls. + earliest_rel fmlls earliest ==> + earliest_rel (list_delete_list l fmlls) earliest +Proof + rw[earliest_rel_def]>> + cheat +QED + +Theorem all_core_list_SORTED: + SORTED $>= inds ∧ + all_core_list fmlls inds [] = SOME inds' ⇒ + SORTED $>= inds' +Proof + rw[]>>drule all_core_list_inds>>rw[]>> + match_mp_tac SORTED_FILTER>> + fs[transitive_def] +QED + Theorem fml_rel_check_cstep_list: fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ + SORTED $>= inds ∧ + earliest_rel fmlls earliest ∧ (∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE) ∧ - check_cstep_list cstep fmlls inds pc = - SOME (fmlls',inds',pc') ⇒ + check_cstep_list cstep fmlls inds earliest pc = + SOME (fmlls',inds',earliest',pc') ⇒ ∃fml'. check_cstep cstep fml pc = SOME (fml', pc') ∧ fml_rel fml' fmlls' ∧ ind_rel fmlls' inds' ∧ + SORTED $>= inds' ∧ + earliest_rel fmlls' earliest' ∧ (∀n. n ≥ pc'.id ⇒ any_el n fmlls' NONE = NONE) ∧ pc.id ≤ pc'.id Proof @@ -1931,6 +2072,12 @@ Proof match_mp_tac ind_rel_rollback_2>> fs[]>> metis_tac[ind_rel_reindex])>> + CONJ_TAC >- + metis_tac[SORTED_sorted_insert,SORTED_reindex]>> + CONJ_TAC >- ( + match_mp_tac earliest_rel_update_resize_update_earliest>> + match_mp_tac fml_rel_fml_rel_earliest_rel>>fs[]>> + match_mp_tac fml_rel_rollback>>rw[]>>fs[])>> simp[rollback_def,any_el_list_delete_list,MEM_MAP,MEM_COUNT_LIST]) >- ( (* Sstep *) gvs[check_cstep_list_def,AllCaseEqs(),check_cstep_def]>> @@ -1945,6 +2092,7 @@ Proof simp[]>> drule_at (Pos last) fml_rel_check_red_list>> disch_then match_mp_tac>> + simp[]>> CONJ_TAC >- ( drule fml_rel_list_delete_list>> disch_then(qspec_then`[n]` mp_tac)>> @@ -1953,6 +2101,10 @@ Proof drule ind_rel_list_delete_list>> disch_then(qspec_then`[n]` mp_tac)>> simp[list_delete_list_def])>> + CONJ_TAC >- ( + drule earliest_rel_list_delete_list>> + disch_then(qspec_then`[n]` mp_tac)>> + simp[list_delete_list_def])>> metis_tac[any_el_list_delete_list,list_delete_list_def]) >- ( (* UncheckedDelete *) gvs[check_cstep_list_def,AllCaseEqs(),check_cstep_def] @@ -1961,6 +2113,8 @@ Proof metis_tac[fml_rel_list_delete_list]>> CONJ_TAC >- metis_tac[ind_rel_list_delete_list]>> + CONJ_TAC >- + metis_tac[earliest_rel_list_delete_list]>> simp[any_el_list_delete_list]) >- ( drule_all fml_rel_all_core>>strip_tac>> @@ -1969,6 +2123,10 @@ Proof metis_tac[fml_rel_list_delete_list]>> CONJ_TAC >- metis_tac[ind_rel_list_delete_list]>> + CONJ_TAC >- + metis_tac[all_core_list_SORTED]>> + CONJ_TAC >- + metis_tac[earliest_rel_list_delete_list]>> simp[any_el_list_delete_list])) >- ( (* Transfer *) gvs[check_cstep_list_def,AllCaseEqs(),check_cstep_def]>> @@ -1977,11 +2135,14 @@ Proof strip_tac>>fs[]>> fs[ind_rel_def]>> rw[]>> + (* TODO should be an easy induction *) + `earliest_rel fmlls' earliest` by cheat>> metis_tac[IS_SOME_EXISTS,option_CLAUSES]) >- ( (* StrengthenToCore *) gvs[check_cstep_list_def,AllCaseEqs(),check_cstep_def]>> Cases_on`reindex F fmlls inds`>> - drule_all ind_rel_reindex + drule_all ind_rel_reindex>> + drule_all SORTED_reindex >- ( drule any_el_core_from_inds>> rw[] @@ -1993,19 +2154,24 @@ Proof `any_el x fmlls NONE = NONE` by metis_tac[IS_SOME_EXISTS,option_CLAUSES]>> simp[]>> - metis_tac[fml_rel_def])>> - fs[ind_rel_def]>> - rw[]>> - metis_tac[IS_SOME_EXISTS,option_CLAUSES]) + metis_tac[fml_rel_def]) + >- ( + fs[ind_rel_def]>> + rw[]>> + metis_tac[IS_SOME_EXISTS,option_CLAUSES]) + >- + cheat) >- fs[]) >- ( (* LoadOrder *) gvs[check_cstep_list_def,AllCaseEqs(),check_cstep_def]>> Cases_on`reindex F fmlls inds`>> drule_all ind_rel_reindex>> + drule_all SORTED_reindex>> drule any_el_core_from_inds>> strip_tac>>fs[]>> strip_tac>>fs[]>> + strip_tac>>fs[]>> CONJ_TAC >- ( simp[fml_rel_def,lookup_map]>> fs[ind_rel_def]>>rw[] @@ -2015,9 +2181,11 @@ Proof metis_tac[IS_SOME_EXISTS,option_CLAUSES]>> simp[]>> metis_tac[fml_rel_def])>> - fs[ind_rel_def]>> - rw[]>> - metis_tac[IS_SOME_EXISTS,option_CLAUSES]) + CONJ_TAC >- ( + fs[ind_rel_def]>> + rw[]>> + metis_tac[IS_SOME_EXISTS,option_CLAUSES])>> + (*easy *) cheat) >- ( (* UnloadOrder *) gvs[check_cstep_list_def,AllCaseEqs(),check_cstep_def]) >- ( (* StoreOrder *) @@ -2034,7 +2202,9 @@ Proof pairarg_tac>>gvs[]>> rw[]>>gvs[] >- metis_tac[fml_rel_update_resize] - >- metis_tac[ind_rel_update_resize_sorted_insert]>> + >- metis_tac[ind_rel_update_resize_sorted_insert] + >- metis_tac[SORTED_sorted_insert] + >- metis_tac[earliest_rel_update_resize_update_earliest]>> simp[any_el_update_resize]) >- ( (* ChangeObj *) fs[check_cstep_def,check_cstep_list_def]>> @@ -2060,12 +2230,14 @@ Proof drule check_subproofs_list_id_upper>> drule check_subproofs_list_mindel>> ntac 3 strip_tac>> - CONJ_TAC >- ( + CONJ_ASM1_TAC >- ( match_mp_tac fml_rel_rollback>>rw[]>>fs[])>> CONJ_TAC >- ( match_mp_tac ind_rel_rollback_2>> simp[]>> metis_tac[ind_rel_reindex])>> + CONJ_TAC >- + metis_tac[fml_rel_fml_rel_earliest_rel]>> simp[any_el_rollback]) >- ( (* CheckObj *) fs[check_cstep_def,check_cstep_list_def] @@ -2073,25 +2245,30 @@ Proof QED Definition check_csteps_list_def: - (check_csteps_list [] fml inds pc = SOME (fml, inds, pc)) ∧ - (check_csteps_list (c::cs) fml inds pc = - case check_cstep_list c fml inds pc of + (check_csteps_list [] fml inds earliest pc = + SOME (fml, inds, earliest, pc)) ∧ + (check_csteps_list (c::cs) fml inds earliest pc = + case check_cstep_list c fml inds earliest pc of NONE => NONE - | SOME(fml', inds', pc') => - check_csteps_list cs fml' inds' pc') + | SOME(fml', inds', earliest', pc') => + check_csteps_list cs fml' inds' earliest' pc') End Theorem fml_rel_check_csteps_list: - ∀csteps fml fmlls inds pc fmlls' inds' pc'. + ∀csteps fml fmlls inds earliest pc fmlls' inds' earliest' pc'. fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ + SORTED $>= inds ∧ + earliest_rel fmlls earliest ∧ (∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE) ∧ - check_csteps_list csteps fmlls inds pc = - SOME (fmlls', inds', pc') ⇒ + check_csteps_list csteps fmlls inds earliest pc = + SOME (fmlls', inds', earliest', pc') ⇒ ∃fml'. check_csteps csteps fml pc = SOME (fml', pc') ∧ fml_rel fml' fmlls' ∧ ind_rel fmlls' inds' ∧ + SORTED $>= inds' ∧ + earliest_rel fmlls' earliest' ∧ (∀n. n ≥ pc'.id ⇒ any_el n fmlls' NONE = NONE) ∧ pc.id ≤ pc'.id Proof @@ -2288,24 +2465,64 @@ Proof gs[EL_REPLICATE] QED +Theorem earliest_rel_FOLDL_update_resize_aux: + ∀xs ls t. + earliest_rel ls t ⇒ + earliest_rel + (FOLDL (λacc (i,v). update_resize acc NONE (SOME (v,b)) i) ls xs) + (FOLDL (λacc (i,v). update_earliest acc i (FST v)) t xs) +Proof + Induct>>rw[]>> + first_x_assum match_mp_tac>> + pairarg_tac>>gvs[]>> + match_mp_tac earliest_rel_update_resize_update_earliest>> + fs[] +QED + +Theorem earliest_rel_FOLDL_update_resize: + earliest_rel + (FOLDL (λacc (i,v). update_resize acc NONE (SOME (v,b)) i) (REPLICATE n NONE) (enumerate k fml)) + (FOLDL (λacc (i,v). update_earliest acc i (FST v)) LN (enumerate k fml)) +Proof + match_mp_tac earliest_rel_FOLDL_update_resize_aux>> + simp[earliest_rel_def,min_opt_def,EL_REPLICATE] +QED + +Theorem SORTED_REVERSE_enumerate: + ∀(ls:'a list) k. + SORTED $>= (REVERSE (MAP FST (enumerate k ls))) +Proof + Induct>>rw[miscTheory.enumerate_def]>> + match_mp_tac SORTED_APPEND_IMP>> + rw[transitive_def]>> + fs[MAP_FST_enumerate,MEM_GENLIST] +QED + Theorem check_csteps_list_concl: check_csteps_list cs (FOLDL (λacc (i,v). update_resize acc NONE (SOME (v,T)) i) (REPLICATE m NONE) (enumerate 1 fml)) (REVERSE (MAP FST (enumerate 1 fml))) + (FOLDL (λacc (i,v). update_earliest acc i (FST v)) LN (enumerate 1 fml)) (init_conf (LENGTH fml + 1) chk obj) = - SOME(fmlls',inds',pc') ∧ + SOME(fmlls',inds',earliest',pc') ∧ check_hconcl_list fml obj fmlls' pc'.obj pc'.bound pc'.dbound hconcl ⇒ sem_concl (set fml) obj (hconcl_concl hconcl) Proof rw[]>> - qmatch_asmsub_abbrev_tac`check_csteps_list cs fmlls inds pc = _`>> + qmatch_asmsub_abbrev_tac`check_csteps_list cs fmlls inds + earliest pc = _`>> `fml_rel (build_fml T 1 fml) fmlls` by simp[Abbr`fmlls`,fml_rel_FOLDL_update_resize]>> `ind_rel fmlls inds` by ( unabbrev_all_tac>> simp[ind_rel_FOLDL_update_resize])>> + `SORTED $>= inds` by + (unabbrev_all_tac>>fs[SORTED_REVERSE_enumerate])>> + `earliest_rel fmlls earliest` by ( + unabbrev_all_tac>> + simp[earliest_rel_FOLDL_update_resize])>> `∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE` by ( rw[Abbr`pc`,Abbr`fmlls`,any_el_ALT,init_conf_def]>> DEP_REWRITE_TAC [FOLDL_update_resize_lookup]>> @@ -2401,19 +2618,25 @@ Theorem check_csteps_list_output: (FOLDL (λacc (i,v). update_resize acc NONE (SOME (v,T)) i) (REPLICATE m NONE) (enumerate 1 fml)) (REVERSE (MAP FST (enumerate 1 fml))) + (FOLDL (λacc (i,v). update_earliest acc i (FST v)) LN (enumerate 1 fml)) (init_conf (LENGTH fml + 1) chk obj) = - SOME(fmlls',inds',pc') ∧ + SOME(fmlls',inds',earliest',pc') ∧ check_output_list fmlls' inds' pc'.obj pc'.bound pc'.dbound pc'.chk fmlt objt output ⇒ sem_output (set fml) obj pc'.bound (set fmlt) objt output Proof rw[]>> - qmatch_asmsub_abbrev_tac`check_csteps_list cs fmlls inds pc = _`>> + qmatch_asmsub_abbrev_tac`check_csteps_list cs fmlls inds earliest pc = _`>> `fml_rel (build_fml T 1 fml) fmlls` by simp[Abbr`fmlls`,fml_rel_FOLDL_update_resize]>> `ind_rel fmlls inds` by ( unabbrev_all_tac>> simp[ind_rel_FOLDL_update_resize])>> + `SORTED $>= inds` by + (unabbrev_all_tac>>fs[SORTED_REVERSE_enumerate])>> + `earliest_rel fmlls earliest` by ( + unabbrev_all_tac>> + simp[earliest_rel_FOLDL_update_resize])>> `∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE` by ( rw[Abbr`pc`,Abbr`fmlls`,any_el_ALT,init_conf_def]>> DEP_REWRITE_TAC [FOLDL_update_resize_lookup]>> From f3f3219476fff57e3b9c78604d5715b4a10db9d8 Mon Sep 17 00:00:00 2001 From: Magnus Myreen Date: Sun, 21 Jan 2024 16:35:54 +0100 Subject: [PATCH 04/38] Remove a cheat --- examples/pseudo_bool/array/npbc_listScript.sml | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 958da406ac..398792fb84 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -1209,7 +1209,11 @@ Proof gvs[reindex_partial_def,AllCaseEqs()]>> drule reindex_partial_aux>> strip_tac>>gvs[]>> - cheat + conj_tac >- (gvs [ind_rel_def,MEM_FILTER]) + \\ DEP_REWRITE_TAC [SORTED_APPEND] + \\ gvs [MEM_FILTER] + \\ rpt $ irule_at Any sortingTheory.SORTED_FILTER \\ gvs [] + \\ gvs [transitive_def] QED Theorem MEM_subst_indexes: From 88bcf09b1e175d800af06e6ea3543d5db0ca68bb Mon Sep 17 00:00:00 2001 From: Magnus Myreen Date: Sun, 21 Jan 2024 18:39:22 +0100 Subject: [PATCH 05/38] Attempt to make progress on a cheat --- .../pseudo_bool/array/npbc_listScript.sml | 30 +++++++++++++++++-- 1 file changed, 28 insertions(+), 2 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 398792fb84..54caf5297a 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -1582,13 +1582,39 @@ Proof gvs[] QED -(* TODO: maybe some induction? *) +Triviality earliest_rel_append_NONE: + earliest_rel fml earliest ⇒ + earliest_rel (fml ++ REPLICATE k NONE) earliest +Proof + gvs [earliest_rel_def] \\ rw [] + \\ Cases_on ‘LENGTH fml ≤ pos’ >- + (simp [EL_APPEND2] + \\ DEP_REWRITE_TAC [EL_REPLICATE] \\ fs [] + \\ Cases_on ‘lookup x earliest’ \\ gvs [min_opt_def]) + \\ gvs [GSYM NOT_LESS,EL_APPEND1] + \\ last_x_assum irule + \\ Cases_on ‘lookup x earliest’ \\ gvs [min_opt_def] +QED + +Triviality earliest_rel_lupdate: + n < LENGTH fml ∧ + earliest_rel fml earliest ⇒ + earliest_rel (LUPDATE (SOME (v,b)) n fml) + (update_earliest earliest n (FST v)) +Proof + gvs [earliest_rel_def] \\ rw [] \\ gvs [EL_LUPDATE] + \\ PairCases_on ‘v’ \\ gvs [] + \\ cheat +QED + Theorem earliest_rel_update_resize_update_earliest: earliest_rel fml earliest ⇒ earliest_rel (update_resize fml NONE (SOME (v,b)) n) (update_earliest earliest n (FST v)) Proof - cheat + gvs [update_resize_def] \\ IF_CASES_TAC \\ strip_tac + \\ irule earliest_rel_lupdate \\ fs [] + \\ irule earliest_rel_append_NONE \\ fs [] QED Theorem opt_update_inds_earliest_rel: From f3d14e036fa317db68d405df926be4dcd05c6de5 Mon Sep 17 00:00:00 2001 From: Magnus Myreen Date: Sun, 21 Jan 2024 19:21:51 +0100 Subject: [PATCH 06/38] Remove cheat --- .../pseudo_bool/array/npbc_listScript.sml | 45 ++++++++++++++++++- 1 file changed, 44 insertions(+), 1 deletion(-) diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 54caf5297a..d9889e11b9 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -1596,6 +1596,28 @@ Proof \\ Cases_on ‘lookup x earliest’ \\ gvs [min_opt_def] QED +Triviality lookup_update_earliest_none: + ∀v0 n earliest x. + lookup x (update_earliest earliest n v0) = NONE ⇒ + ¬MEM x (MAP SND v0) ∧ lookup x earliest = NONE +Proof + Induct \\ gvs [update_earliest_def,FORALL_PROD] + \\ rw [] \\ first_x_assum drule \\ fs [] + \\ gvs [lookup_insert,AllCaseEqs()] +QED + +Triviality lookup_update_earliest_some: + ∀v0 n earliest x k. + lookup x (update_earliest earliest n v0) = SOME k ∧ n < k ⇒ + ¬MEM x (MAP SND v0) ∧ lookup x earliest = SOME k +Proof + Induct \\ gvs [update_earliest_def,FORALL_PROD] + \\ rw [] \\ first_x_assum drule \\ fs [] + \\ gvs [lookup_insert,AllCaseEqs()] + \\ Cases_on ‘lookup p_2 earliest’ \\ gvs [min_opt_def] + \\ ‘MIN x' n ≠ k’ by gvs [MIN_DEF] \\ gvs [] +QED + Triviality earliest_rel_lupdate: n < LENGTH fml ∧ earliest_rel fml earliest ⇒ @@ -1604,7 +1626,28 @@ Triviality earliest_rel_lupdate: Proof gvs [earliest_rel_def] \\ rw [] \\ gvs [EL_LUPDATE] \\ PairCases_on ‘v’ \\ gvs [] - \\ cheat + \\ IF_CASES_TAC \\ gvs [] + >- + (Cases_on ‘lookup x (update_earliest earliest n v0)’ \\ gvs [min_opt_def] + \\ imp_res_tac lookup_update_earliest_none + \\ imp_res_tac lookup_update_earliest_some) + \\ first_x_assum irule + \\ Cases_on ‘lookup x (update_earliest earliest n v0)’ \\ gvs [] + >- (imp_res_tac lookup_update_earliest_none \\ gvs []) + \\ gvs [min_opt_def] + \\ Cases_on ‘lookup x earliest’ \\ gvs [] + \\ irule LESS_LESS_EQ_TRANS + \\ last_x_assum $ irule_at $ Pos hd + \\ rename [‘a ≤ b’] + \\ pop_assum mp_tac + \\ pop_assum mp_tac + \\ qid_spec_tac ‘earliest’ + \\ qid_spec_tac ‘b’ + \\ qid_spec_tac ‘a’ + \\ Induct_on ‘v0’ \\ gvs [update_earliest_def,FORALL_PROD] + \\ rw [] \\ first_x_assum drule + \\ gvs [lookup_insert] \\ rw [] + \\ gvs [min_opt_def] QED Theorem earliest_rel_update_resize_update_earliest: From 69aa95ce9f2d5c1ded7121063b8579c6301fc4f1 Mon Sep 17 00:00:00 2001 From: Magnus Myreen Date: Sun, 21 Jan 2024 19:30:17 +0100 Subject: [PATCH 07/38] Remove cheat --- examples/pseudo_bool/array/npbc_listScript.sml | 10 +++++++--- 1 file changed, 7 insertions(+), 3 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index d9889e11b9..0a0b5a9227 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -2066,14 +2066,18 @@ Proof simp[rollback_def,any_el_list_delete_list,MEM_MAP,MEM_COUNT_LIST] QED -(* should be easy *) Theorem earliest_rel_list_delete_list: ∀l fmlls. earliest_rel fmlls earliest ==> earliest_rel (list_delete_list l fmlls) earliest Proof - rw[earliest_rel_def]>> - cheat + Induct \\ gvs [list_delete_list_def] \\ rw [] + \\ last_x_assum irule + \\ gvs [earliest_rel_def] + \\ ‘LENGTH (delete_list h fmlls) = LENGTH fmlls’ by rw [delete_list_def] + \\ gvs [] \\ rw [] + \\ last_x_assum drule + \\ rw [delete_list_def] \\ rw [EL_LUPDATE] QED Theorem all_core_list_SORTED: From 98ae4cbc5cd34ae4092b6e89a7a468d1ba199aeb Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Wed, 24 Jan 2024 12:54:23 +0800 Subject: [PATCH 08/38] pass earliest through --- .../array/npbc_arrayProgScript.sml | 289 +++++++++++++----- .../array/npbc_parseProgScript.sml | 125 +++++--- 2 files changed, 300 insertions(+), 114 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_arrayProgScript.sml b/examples/pseudo_bool/array/npbc_arrayProgScript.sml index 7f0190bc77..ed3bd2aa14 100644 --- a/examples/pseudo_bool/array/npbc_arrayProgScript.sml +++ b/examples/pseudo_bool/array/npbc_arrayProgScript.sml @@ -1023,6 +1023,113 @@ Proof metis_tac[LIST_TYPE_def] QED +val reindex_partial_aux_arr = process_topdecs ` + fun reindex_partial_aux_arr b fml mini ls iacc vacc = + case ls of + [] => (List.rev iacc, (vacc, [])) + | (i::is) => + if i < mini then (List.rev iacc, (vacc, ls)) + else + case Array.lookup fml None i of + None => reindex_partial_aux_arr b fml mini is iacc vacc + | Some (v,b') => + reindex_partial_aux_arr b fml mini is (i::iacc) + (mk_vacc b b' v vacc)` |> append_prog; + +val reindex_partial_arr = process_topdecs ` + fun reindex_partial_arr b fml mini is = + case mini of None => ([], ([], is)) + | Some mini => + reindex_partial_aux_arr b fml mini is [] []` + |> append_prog; + +Theorem reindex_partial_aux_arr_spec: + ∀inds indsv b bv fmlls fmlv iacc iaccv vacc vaccv. + BOOL b bv ∧ + LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ + NUM mini miniv ∧ + (LIST_TYPE NUM) inds indsv ∧ + (LIST_TYPE NUM) iacc iaccv ∧ + (LIST_TYPE constraint_TYPE) vacc vaccv + ⇒ + app (p : 'ffi ffi_proj) + ^(fetch_v "reindex_partial_aux_arr" (get_ml_prog_state())) + [bv; fmlv; miniv; indsv; iaccv; vaccv] + (ARRAY fmlv fmllsv) + (POSTv v. + ARRAY fmlv fmllsv * + &(PAIR_TYPE (LIST_TYPE NUM) + (PAIR_TYPE (LIST_TYPE constraint_TYPE) + (LIST_TYPE NUM)) + (reindex_partial_aux b fmlls mini inds iacc vacc) v)) +Proof + Induct>> + xcf"reindex_partial_aux_arr"(get_ml_prog_state ())>> + fs[LIST_TYPE_def] + >- ( + xmatch>> + rpt xlet_autop>> + xcon>>xsimpl>> + simp[reindex_partial_aux_def,LIST_TYPE_def,PAIR_TYPE_def])>> + xmatch>> + xlet_autop>> + xif + >- ( + rpt xlet_autop>> + xcon>> + simp[reindex_partial_aux_def,PAIR_TYPE_def,LIST_TYPE_def]>> + xsimpl)>> + rpt xlet_autop>> + xlet_auto>> + `OPTION_TYPE bconstraint_TYPE + (any_el h fmlls NONE) v'` by ( + rw[any_el_ALT]>> + fs[LIST_REL_EL_EQN,OPTION_TYPE_def])>> + rw[]>> + simp[reindex_partial_aux_def]>> + TOP_CASE_TAC>>fs[OPTION_TYPE_def] + >- ( + xmatch>> + xapp>>simp[])>> + TOP_CASE_TAC>>fs[PAIR_TYPE_def]>> + xmatch>> + rpt xlet_autop>> + xapp>> + fs[LIST_TYPE_def,mk_vacc_def] +QED + +Theorem reindex_partial_arr_spec: + ∀inds indsv fmlls fmlv. + BOOL b bv ∧ + OPTION_TYPE NUM mini miniv ∧ + LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ + (LIST_TYPE NUM) inds indsv + ⇒ + app (p : 'ffi ffi_proj) + ^(fetch_v "reindex_partial_arr" (get_ml_prog_state())) + [bv; fmlv; miniv; indsv] + (ARRAY fmlv fmllsv) + (POSTv v. + ARRAY fmlv fmllsv * + &(PAIR_TYPE (LIST_TYPE NUM) + (PAIR_TYPE (LIST_TYPE constraint_TYPE) (LIST_TYPE NUM)) + (reindex_partial b fmlls mini inds) v)) +Proof + xcf"reindex_partial_arr"(get_ml_prog_state ())>> + Cases_on`mini`>> + gvs[OPTION_TYPE_def]>> + xmatch + >- ( + rpt xlet_autop>> + xcon>>xsimpl>> + simp[reindex_partial_def,PAIR_TYPE_def,LIST_TYPE_def])>> + rpt xlet_autop>> + xapp>> + xsimpl>> + simp[reindex_partial_def]>> + metis_tac[LIST_TYPE_def] +QED + val res = translate is_Pos_def; val res = translate subst_aux_def; val res = translate partition_def; @@ -1051,7 +1158,7 @@ val extract_clauses_arr = process_topdecs` else raise Fail (format_failure lno ("invalid #proofgoal id: " ^ Int.toString u)) ` |> append_prog; -Overload "subst_TYPE" = ``(VECTOR_TYPE (OPTION_TYPE (SUM_TYPE BOOL (PBC_LIT_TYPE NUM))))`` +Overload "subst_TYPE" = ``SUM_TYPE (PAIR_TYPE NUM (SUM_TYPE BOOL (PBC_LIT_TYPE NUM))) (VECTOR_TYPE (OPTION_TYPE (SUM_TYPE BOOL (PBC_LIT_TYPE NUM))))`` Overload "pfs_TYPE" = ``LIST_TYPE (PAIR_TYPE (OPTION_TYPE (PAIR_TYPE (SUM_TYPE NUM NUM) NUM)) (LIST_TYPE NPBC_CHECK_LSTEP_TYPE))`` @@ -1722,9 +1829,12 @@ val format_failure_2_def = Define` val r = translate format_failure_2_def; Theorem vec_eq_nil_thm: - v = Vector [] ⇔ length v = 0 + v = INR (Vector []) ⇔ + case v of INL _ => F + | INR v => length v = 0 Proof - EVAL_TAC>>Cases_on`v`>>fs[mlvectorTheory.length_def] + Cases_on`v`>>EVAL_TAC>> + Cases_on`y`>>fs[mlvectorTheory.length_def] QED val r = translate (red_fast_def |> SIMP_RULE std_ss [vec_eq_nil_thm]); @@ -1743,13 +1853,25 @@ val check_red_arr_fast = process_topdecs` else raise Fail (format_failure lno ("did not derive contradiction from index:" ^ Int.toString cid)) end` |> append_prog; +val res = translate get_earliest_def; + +val r = translate spt_to_vecTheory.prepend_def; +val r = translate spt_to_vecTheory.to_flat_def; + +val r = translate spt_to_vecTheory.spt_to_vec_def; +val res = translate fromAList_def; +val res = translate npbc_checkTheory.mk_subst_def; + val check_red_arr = process_topdecs` - fun check_red_arr lno ord obj b tcb fml inds id c s pfs idopt = + fun check_red_arr lno ord obj b tcb fml inds id c s pfs idopt earliest = + let val s = mk_subst s in case red_fast s idopt pfs of None => ( - let val bortcb = b orelse tcb in - case reindex_arr bortcb fml inds of (rinds,fmlls) => + let val bortcb = b orelse tcb + val mini = get_earliest earliest s in + case reindex_partial_arr bortcb fml mini inds of + (rinds,(fmlls,rest)) => let val nc = not_1 c val rsubs = do_rso ord s c obj @@ -1762,20 +1884,21 @@ val check_red_arr = process_topdecs` let val u = rollback_arr fml' id id' val goals = subst_indexes_arr s bortcb fml' rinds in case red_cond_check fmlls nc pfs rsubs goals - of None => (fml', (rinds,id')) + of None => (fml', (rinds @ rest,id')) | Some err => raise Fail (format_failure_2 lno ("Redundancy subproofs did not cover all subgoals. Info: " ^ err ^ ".") (print_subproofs_err rsubs goals)) end | Some cid => if check_contradiction_fml_arr b fml' cid then let val u = rollback_arr fml' id id' in - (fml', (rinds, id')) + (fml', (rinds @ rest, id')) end else raise Fail (format_failure lno ("did not derive contradiction from index:" ^ Int.toString cid))) end end) | Some (pf,cid) => - check_red_arr_fast lno b fml inds id c pf cid` |> append_prog; + check_red_arr_fast lno b fml inds id c pf cid + end` |> append_prog; (* Overloads all the _TYPEs that we will reuse *) Overload "spo_TYPE" = ``PAIR_TYPE @@ -1854,6 +1977,8 @@ Proof metis_tac[ARRAY_refl,Fail_exn_def] QED +Overload "subst_raw_TYPE" = ``LIST_TYPE (PAIR_TYPE NUM (SUM_TYPE BOOL (PBC_LIT_TYPE NUM)))`` + Theorem check_red_arr_spec: NUM lno lnov ∧ ord_TYPE ord ordv ∧ @@ -1864,14 +1989,15 @@ Theorem check_red_arr_spec: (LIST_TYPE NUM) inds indsv ∧ NUM id idv ∧ constraint_TYPE c cv ∧ - subst_TYPE s sv ∧ + subst_raw_TYPE s sv ∧ pfs_TYPE pfs pfsv ∧ - OPTION_TYPE NUM idopt idoptv + OPTION_TYPE NUM idopt idoptv ∧ + SPTREE_SPT_TYPE NUM earliest earliestv ⇒ app (p : 'ffi ffi_proj) ^(fetch_v "check_red_arr" (get_ml_prog_state())) [lnov; ordv; objv; bv; tcbv; fmlv; indsv; idv; - cv; sv; pfsv; idoptv] + cv; sv; pfsv; idoptv; earliestv] (ARRAY fmlv fmllsv) (POSTve (λv. @@ -1879,7 +2005,7 @@ Theorem check_red_arr_spec: ARRAY fmlv' fmllsv' * &( case check_red_list ord obj b tcb fmlls inds id - c s pfs idopt of NONE => F + c s pfs idopt earliest of NONE => F | SOME res => PAIR_TYPE (λl v. LIST_REL (OPTION_TYPE bconstraint_TYPE) l fmllsv' ∧ @@ -1890,13 +2016,17 @@ Theorem check_red_arr_spec: ARRAY fmlv' fmllsv' * & (Fail_exn e ∧ check_red_list ord obj b tcb fmlls inds id - c s pfs idopt = NONE))) + c s pfs idopt earliest = NONE))) Proof rw[]>> xcf "check_red_arr" (get_ml_prog_state ())>> + xlet_auto + >- ( + xsimpl>> + simp (eq_lemmas()))>> xlet_autop>> rw[check_red_list_def]>> - reverse (Cases_on`red_fast s idopt pfs`) + reverse (Cases_on`red_fast (mk_subst s) idopt pfs`) >- ( simp[]>>Cases_on`x`>>fs[OPTION_TYPE_def,PAIR_TYPE_def]>> xmatch>> @@ -1910,22 +2040,23 @@ Proof xvar>>xsimpl)>> pairarg_tac>>gs[]>> xlet_autop>> + xlet_autop>> fs[PAIR_TYPE_def]>> xmatch>> rpt xlet_autop>> - qmatch_asmsub_abbrev_tac`VECTOR_TYPE _ s _`>> + qmatch_asmsub_abbrev_tac`subst_TYPE ss _`>> qmatch_goalsub_abbrev_tac`ARRAY aa vv`>> xlet`(POSTve (λv. ARRAY aa vv * - &(case extract_clauses_list s b fmlls - (do_rso ord s c obj) pfs [] of + &(case extract_clauses_list ss b fmlls + (do_rso ord ss c obj) pfs [] of NONE => F | SOME res => LIST_TYPE (PAIR_TYPE (OPTION_TYPE (PAIR_TYPE (LIST_TYPE constraint_TYPE) NUM)) (LIST_TYPE NPBC_CHECK_LSTEP_TYPE)) res v)) (λe. ARRAY aa vv * & (Fail_exn e ∧ - extract_clauses_list s b fmlls (do_rso ord s c obj) pfs [] = NONE)))` + extract_clauses_list ss b fmlls (do_rso ord ss c obj) pfs [] = NONE)))` >- ( xapp>>xsimpl>> fs[LIST_TYPE_def]>> @@ -1968,6 +2099,7 @@ Proof fs[red_cond_check_def]>> pairarg_tac>>fs[]>> xlet_autop>> + xlet_autop>> xcon>>xsimpl>> simp[PAIR_TYPE_def]>> xsimpl) @@ -1993,25 +2125,31 @@ Proof xsimpl QED +val res = translate min_opt_def; +val res = translate update_earliest_def; + val check_sstep_arr = process_topdecs` - fun check_sstep_arr lno step ord obj tcb fml inds id = + fun check_sstep_arr lno step ord obj tcb fml inds id earliest = case step of Lstep lstep => (case check_lstep_arr lno lstep False fml 0 id of (rfml,(c,id')) => (case c of - None => (rfml,(inds,id')) + None => (rfml,(inds,(earliest,id'))) | Some cc => (Array.updateResize rfml None id' (Some cc), (sorted_insert id' inds, - id'+1)) )) + (update_earliest earliest id' (fst (fst cc)), + id'+1))) )) | Red c s pfs idopt => (case check_red_arr lno ord obj False tcb - fml inds id c s pfs idopt of + fml inds id c s pfs idopt earliest of (fml',(rinds,id')) => (Array.updateResize fml' None id' (Some (c,tcb)), - (sorted_insert id' rinds,(id'+1)))) + (sorted_insert id' rinds, + (update_earliest earliest id' (fst c), + id'+1)))) ` |> append_prog val NPBC_CHECK_SSTEP_TYPE_def = theorem "NPBC_CHECK_SSTEP_TYPE_def"; @@ -2024,11 +2162,12 @@ Theorem check_sstep_arr_spec: BOOL tcb tcbv ∧ LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ (LIST_TYPE NUM) inds indsv ∧ - NUM id idv + NUM id idv ∧ + SPTREE_SPT_TYPE NUM earliest earliestv ⇒ app (p : 'ffi ffi_proj) ^(fetch_v "check_sstep_arr" (get_ml_prog_state())) - [lnov; stepv; ordv; objv; tcbv; fmlv; indsv; idv] + [lnov; stepv; ordv; objv; tcbv; fmlv; indsv; idv; earliestv] (ARRAY fmlv fmllsv) (POSTve (λv. @@ -2036,18 +2175,20 @@ Theorem check_sstep_arr_spec: ARRAY fmlv' fmllsv' * &( case check_sstep_list step ord obj tcb - fmlls inds id of NONE => F + fmlls inds id earliest of NONE => F | SOME res => PAIR_TYPE (λl v. LIST_REL (OPTION_TYPE bconstraint_TYPE) l fmllsv' ∧ - v = fmlv') (PAIR_TYPE (LIST_TYPE NUM) NUM) res v - )) + v = fmlv') + (PAIR_TYPE (LIST_TYPE NUM) + (PAIR_TYPE (SPTREE_SPT_TYPE NUM) NUM)) res v) + ) (λe. SEP_EXISTS fmlv' fmllsv'. ARRAY fmlv' fmllsv' * & (Fail_exn e ∧ check_sstep_list step ord obj tcb - fmlls inds id = NONE))) + fmlls inds id earliest = NONE))) Proof rw[]>> xcf "check_sstep_arr" (get_ml_prog_state ())>> @@ -2089,8 +2230,9 @@ Proof fs[opt_update_inds_def,OPTION_TYPE_def]>> xmatch >- ( - xlet_autop>>xcon>>xsimpl>> - simp[PAIR_TYPE_def]>> + rpt xlet_autop>> + xcon>>xsimpl>> + simp[PAIR_TYPE_def,PULL_EXISTS]>> metis_tac[ARRAY_refl])>> rpt xlet_autop>> xlet_auto>> @@ -2105,14 +2247,14 @@ Proof xmatch>> xlet_autop>> rename1`check_red_list ord obj F tcb fmlls inds id - c s pfs idopt`>> + c s pfs idopt earliest`>> xlet`POSTve (λv. SEP_EXISTS fmlv' fmllsv'. ARRAY fmlv' fmllsv' * &( case check_red_list ord obj F tcb fmlls inds id - c s pfs idopt of NONE => F + c s pfs idopt earliest of NONE => F | SOME res => PAIR_TYPE (λl v. LIST_REL (OPTION_TYPE bconstraint_TYPE) l fmllsv' ∧ @@ -2123,7 +2265,7 @@ Proof ARRAY fmlv' fmllsv' * & (Fail_exn e ∧ check_red_list ord obj F tcb fmlls inds id - c s pfs idopt = NONE))` + c s pfs idopt earliest = NONE))` >- ( xapp >> xsimpl>> CONJ_TAC >- EVAL_TAC>> @@ -2897,22 +3039,23 @@ val res = translate npbc_checkTheory.eq_obj_def; val res = translate npbc_checkTheory.check_eq_obj_def; val check_cstep_arr = process_topdecs` - fun check_cstep_arr lno cstep fml inds pc = + fun check_cstep_arr lno cstep fml inds earliest pc = case cstep of Dom c s pfs idopt => ( case get_ord pc of None => raise Fail (format_failure lno "no order loaded for dominance step") | Some spo => case check_dom_arr lno spo (get_obj pc) - fml inds (get_id pc) c s pfs idopt of (fml',(rinds,id')) => + fml inds (get_id pc) c (mk_subst s) pfs idopt of + (fml',(rinds,id')) => (Array.updateResize fml' None id' (Some (c,get_tcb pc)), (sorted_insert id' rinds, - (set_id pc (id'+1))))) + (update_earliest earliest id' (fst c), set_id pc (id'+1))))) | Sstep sstep => ( case check_sstep_arr lno sstep (get_ord pc) (get_obj pc) - (get_tcb pc) fml inds (get_id pc) of - (fml',(inds',id')) => - (fml',(inds',set_id pc id'))) + (get_tcb pc) fml inds (get_id pc) earliest of + (fml',(inds',(earliest',id'))) => + (fml',(inds',(earliest',set_id pc id')))) | Checkeddelete n s pfs idopt => ( if check_tcb_idopt_pc pc idopt then @@ -2923,35 +3066,35 @@ val check_cstep_arr = process_topdecs` case check_red_arr lno (get_ord pc) (get_obj pc) True (get_tcb pc) fml inds (get_id pc) - c s pfs idopt of (fml',(inds',id')) => - (fml',(inds',set_id pc id')))) + c s pfs idopt earliest of (fml',(inds',id')) => + (fml',(inds',(earliest,set_id pc id'))))) else raise Fail (format_failure lno "invalid proof state for checked deletion")) | Uncheckeddelete ls => ( if check_tcb_ord pc then (list_delete_arr ls fml; - (fml, (inds, set_chk pc False))) + (fml, (inds, (earliest, set_chk pc False)))) else case all_core_arr fml inds [] of None => raise Fail (format_failure lno "not all constraints in core") | Some inds' => (list_delete_arr ls fml; - (fml, (inds', set_chk pc False)))) + (fml, (inds', (earliest, set_chk pc False))))) | Transfer ls => ( let val fml' = core_from_inds_arr lno fml ls in - (fml', (inds, pc)) + (fml', (inds, (earliest, pc))) end) | Strengthentocore b => ( let val inds' = fst (reindex_arr False fml inds) in if b then ( let val fml' = core_from_inds_arr lno fml inds' in - (fml', (inds', set_tcb pc b)) + (fml', (inds', (earliest, set_tcb pc b))) end) else - (fml,(inds',set_tcb pc b)) + (fml,(inds', (earliest, set_tcb pc b))) end) | Loadorder nn xs => let val inds' = fst (reindex_arr False fml inds) in @@ -2961,7 +3104,7 @@ val check_cstep_arr = process_topdecs` | Some ord' => if List.length xs = List.length (fst (snd ord')) then let val fml' = core_from_inds_arr lno fml inds' in - (fml', (inds',set_ord pc (Some (ord',xs)))) + (fml', (inds', (earliest, set_ord pc (Some (ord',xs))))) end else raise Fail @@ -2971,11 +3114,11 @@ val check_cstep_arr = process_topdecs` (case get_ord pc of None => raise Fail (format_failure lno ("no order loaded")) | Some spo => - (fml,(inds, set_ord pc None))) + (fml,(inds,(earliest,set_ord pc None)))) | Storeorder name spo ws pfsr pfst => if check_storeorder spo ws pfst pfsr then - (fml, (inds, set_orders pc ((name,spo)::get_orders pc))) + (fml, (inds, (earliest, set_orders pc ((name,spo)::get_orders pc)))) else raise Fail (format_failure lno ("invalid order definition")) | Obj w mi bopt => ( @@ -2995,20 +3138,21 @@ val check_cstep_arr = process_topdecs` val c = model_improving obj new in (Array.updateResize fml None id (Some (c,True)), (sorted_insert id inds, - (obj_update pc (id+1) bound' dbound'))) + (update_earliest earliest id (fst c), + obj_update pc (id+1) bound' dbound'))) end else - (fml, (inds, (obj_update pc (get_id pc) bound' dbound'))) + (fml, (inds, (earliest, obj_update pc (get_id pc) bound' dbound'))) end ) | Changeobj b fc' pfs => (case check_change_obj_arr lno b fml (get_id pc) (get_obj pc) fc' pfs of (fml',(fc',id')) => - (fml', (inds, (change_obj_update pc id' fc')))) + (fml', (inds, (earliest, change_obj_update pc id' fc')))) | Checkobj fc' => if check_eq_obj (get_obj pc) fc' - then (fml, (inds, pc)) + then (fml, (inds, (earliest, pc))) else raise Fail (format_failure lno (err_obj_check_string (get_obj pc) fc')) @@ -3041,32 +3185,34 @@ Theorem check_cstep_arr_spec: NPBC_CHECK_CSTEP_TYPE cstep cstepv ∧ LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ (LIST_TYPE NUM) inds indsv ∧ - NPBC_CHECK_PROOF_CONF_TYPE pc pcv + NPBC_CHECK_PROOF_CONF_TYPE pc pcv ∧ + SPTREE_SPT_TYPE NUM earliest earliestv ⇒ app (p : 'ffi ffi_proj) ^(fetch_v "check_cstep_arr" (get_ml_prog_state())) - [lnov; cstepv; fmlv; indsv; pcv] + [lnov; cstepv; fmlv; indsv; earliestv; pcv] (ARRAY fmlv fmllsv) (POSTve (λv. SEP_EXISTS fmlv' fmllsv'. ARRAY fmlv' fmllsv' * &( - case check_cstep_list cstep fmlls inds pc of + case check_cstep_list cstep fmlls inds earliest pc of NONE => F | SOME res => PAIR_TYPE (λl v. LIST_REL (OPTION_TYPE bconstraint_TYPE) l fmllsv' ∧ v = fmlv') (PAIR_TYPE (LIST_TYPE NUM) - NPBC_CHECK_PROOF_CONF_TYPE) + (PAIR_TYPE (SPTREE_SPT_TYPE NUM) + NPBC_CHECK_PROOF_CONF_TYPE)) res v )) (λe. SEP_EXISTS fmlv' fmllsv'. ARRAY fmlv' fmllsv' * & (Fail_exn e ∧ - check_cstep_list cstep fmlls inds pc = NONE))) + check_cstep_list cstep fmlls inds earliest pc = NONE))) Proof rw[]>> xcf "check_cstep_arr" (get_ml_prog_state ())>> @@ -3083,6 +3229,9 @@ Proof metis_tac[Fail_exn_def,ARRAY_refl])>> pairarg_tac>>fs[]>> rpt xlet_autop>> + xlet_auto >- + (xsimpl>> simp (eq_lemmas()))>> + rpt xlet_autop>> fs[get_id_def,get_obj_def] >- ( xsimpl>> @@ -3149,14 +3298,14 @@ Proof rpt xlet_autop>> fs[get_id_def,get_tcb_def,get_obj_def,get_ord_def]>> rename1`check_red_list ord obj T pc.tcb (delete_list n fmlls) - inds pc.id c s pfs idopt`>> + inds pc.id c s pfs idopt earliest`>> xlet`POSTve (λv. SEP_EXISTS fmlv' fmllsv'. ARRAY fmlv' fmllsv' * &( case check_red_list ord obj T pc.tcb (delete_list n fmlls) - inds pc.id c s pfs idopt of NONE => F + inds pc.id c s pfs idopt earliest of NONE => F | SOME res => PAIR_TYPE (λl v. LIST_REL (OPTION_TYPE bconstraint_TYPE) l fmllsv' ∧ @@ -3167,7 +3316,7 @@ Proof ARRAY fmlv' fmllsv' * & (Fail_exn e ∧ check_red_list ord obj T pc.tcb (delete_list n fmlls) - inds pc.id c s pfs idopt = NONE))` + inds pc.id c s pfs idopt earliest = NONE))` >- ( xapp>>xsimpl>> CONJ_TAC >- EVAL_TAC>> @@ -3196,7 +3345,7 @@ Proof qexists_tac`F`>> qexists_tac`pc`>>simp[]>> EVAL_TAC)>> - xlet_autop>> + rpt xlet_autop>> xcon>>xsimpl>> fs[]>> fs[PAIR_TYPE_def,set_chk_def]>> @@ -3221,7 +3370,7 @@ Proof qexists_tac`F`>> qexists_tac`pc`>>simp[]>> EVAL_TAC)>> - xlet_autop>> + rpt xlet_autop>> xcon>>xsimpl >> fs[PAIR_TYPE_def]>> asm_exists_tac>>simp[]>> @@ -3233,7 +3382,7 @@ Proof >- ( xsimpl>> metis_tac[ARRAY_refl])>> - xlet_autop>> + rpt xlet_autop>> xcon>>every_case_tac>>fs[]>>xsimpl>> simp[PAIR_TYPE_def]>> metis_tac[ARRAY_refl]) @@ -3295,7 +3444,7 @@ Proof xapp>>xsimpl>> asm_exists_tac>>simp[]>> qexists_tac`SOME (x,l)`>>simp[OPTION_TYPE_def,PAIR_TYPE_def])>> - xlet_autop>> + rpt xlet_autop>> xcon>>xsimpl>> every_case_tac>> gvs[AllCaseEqs(),PAIR_TYPE_def,OPTION_TYPE_def,set_ord_def]>> @@ -3316,7 +3465,7 @@ Proof xapp>>xsimpl>> asm_exists_tac>>simp[]>> qexists_tac`NONE`>>simp[OPTION_TYPE_def,PAIR_TYPE_def])>> - xlet_autop>> + rpt xlet_autop>> xcon>>xsimpl>> fs[PAIR_TYPE_def,OPTION_TYPE_def,set_ord_def]>> metis_tac[ARRAY_refl]) @@ -3340,7 +3489,7 @@ Proof asm_exists_tac>> qexists_tac`(m,p')::pc.orders`>>xsimpl>> fs[LIST_TYPE_def,PAIR_TYPE_def,get_orders_def])>> - xlet_autop>> + rpt xlet_autop>> xcon>>xsimpl>> fs[PAIR_TYPE_def,OPTION_TYPE_def,LIST_TYPE_def,set_orders_def]>> asm_exists_tac>>simp[]>> @@ -3404,7 +3553,7 @@ Proof fs[get_obj_def]>> xif >- ( - xlet_autop>> + rpt xlet_autop>> xcon>>xsimpl>> simp[PAIR_TYPE_def]>> metis_tac[ARRAY_refl])>> diff --git a/examples/pseudo_bool/array/npbc_parseProgScript.sml b/examples/pseudo_bool/array/npbc_parseProgScript.sml index 5d88050084..eb6bf6f4a6 100644 --- a/examples/pseudo_bool/array/npbc_parseProgScript.sml +++ b/examples/pseudo_bool/array/npbc_parseProgScript.sml @@ -64,10 +64,7 @@ val parse_cutting_side = Q.prove( val r = translate parse_var_def; val r = translate parse_subst_aux_def; -val r = translate spt_to_vecTheory.prepend_def; -val r = translate spt_to_vecTheory.to_flat_def; -val r = translate spt_to_vecTheory.spt_to_vec_def; val r = translate parse_subst_def; val r = translate pbcTheory.lit_var_def; @@ -188,11 +185,11 @@ val int_start_side = Q.prove( val _ = translate tokenize_fast_def; -Definition not_is_empty_vec_def: - not_is_empty_vec v ⇔ length v ≠ 0 +Definition not_is_empty_def: + not_is_empty v ⇔ v ≠ [] End -val _ = translate not_is_empty_vec_def; +val _ = translate not_is_empty_def; val parse_lsteps_aux = process_topdecs` fun parse_lsteps_aux f_ns fd lno acc = @@ -204,7 +201,7 @@ val parse_lsteps_aux = process_topdecs` | Some (Inl step,f_ns') => parse_lsteps_aux f_ns' fd (lno+1) (step::acc) | Some (Inr (c,s),f_ns') => - if not_is_empty_vec s then + if not_is_empty s then raise Fail (format_failure (lno+1) "only contradiction steps allowed in nested proof steps") else (case parse_lsteps_aux f_ns' fd (lno+1) [] of @@ -275,9 +272,9 @@ Proof xsimpl QED -Theorem not_is_empty_vec_eq: - not_is_empty_vec v ⇔ - ¬is_empty_vec v +Theorem not_is_empty_eq: + not_is_empty v ⇔ + ¬is_empty v Proof EVAL_TAC>> Cases_on`v`>>fs[]>> @@ -398,9 +395,11 @@ Proof simp[Once PAIR_TYPE_def]>> strip_tac>> xmatch>> - xlet_autop>> - rename1`is_empty_vec tt`>> - reverse (Cases_on`is_empty_vec tt`>>fs[not_is_empty_vec_eq]) + xlet_auto >- ( + xsimpl>> + simp (eq_lemmas()))>> + rename1`is_empty tt`>> + reverse (Cases_on`is_empty tt`>>fs[not_is_empty_eq]) >- ( xif>>asm_exists_tac>>xsimpl>> rpt xlet_autop>> @@ -679,13 +678,14 @@ Proof metis_tac[STDIO_INSTREAM_LINES_refl_gc]) QED -Theorem is_empty_vec_thm: - is_empty_vec v ⇔ length v = 0 +Theorem is_empty_thm: + is_empty v ⇔ + case v of [] => T | _ => F Proof EVAL_TAC>>Cases_on`v`>>fs[mlvectorTheory.length_def] QED -val res = translate is_empty_vec_thm; +val res = translate is_empty_thm; val res = translate reduce_pf_def; @@ -2037,7 +2037,6 @@ val res = translate parse_obj_term_npbc_def; val res = translate parse_strengthen_def; - val res = translate parse_b_obj_term_npbc_def; val res = translate parse_cstep_head_def; @@ -2313,15 +2312,15 @@ QED (* returns the necessary information to check the output and conclusion sections *) val check_unsat'' = process_topdecs ` - fun check_unsat'' fns fd lno fml inds pc = + fun check_unsat'' fns fd lno fml inds earliest pc = case parse_cstep fns fd lno of (Inl s, (fns', lno')) => (lno', (s, (fns', (fml, (inds, pc))))) | (Inr cstep, (fns', lno')) => - (case check_cstep_arr lno cstep fml inds pc of - (fml', (inds', pc')) => - check_unsat'' fns' fd lno' fml' inds' pc')` |> append_prog + (case check_cstep_arr lno cstep fml inds earliest pc of + (fml', (inds', (earliest', pc'))) => + check_unsat'' fns' fd lno' fml' inds' earliest' pc')` |> append_prog Theorem parse_sstep_LENGTH: ∀f ss res f' ss'. @@ -2366,15 +2365,15 @@ QED returning the last encountered state *) Definition parse_and_run_def: parse_and_run fns ss - fml inds pc = + fml inds earliest pc = case parse_cstep fns ss of NONE => NONE | SOME (INL s, fns', rest) => SOME (rest, s, fns', fml, inds, pc) | SOME (INR cstep, fns', rest) => - (case check_cstep_list cstep fml inds pc of - SOME (fml', inds', pc') => - parse_and_run fns' rest fml' inds' pc' + (case check_cstep_list cstep fml inds earliest pc of + SOME (fml', inds', earliest', pc') => + parse_and_run fns' rest fml' inds' earliest' pc' | res => NONE) Termination WF_REL_TAC `measure (LENGTH o FST o SND)`>> @@ -2413,18 +2412,19 @@ Proof QED Theorem check_unsat''_spec: - ∀fns ss fmlls inds pc - fnsv lno lnov fmllsv indsv pcv lines fs fmlv. + ∀fns ss fmlls inds earliest pc + fnsv lno lnov fmllsv indsv pcv lines fs fmlv earliestv. fns_TYPE a fns fnsv ∧ NUM lno lnov ∧ LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ (LIST_TYPE NUM) inds indsv ∧ NPBC_CHECK_PROOF_CONF_TYPE pc pcv ∧ + SPTREE_SPT_TYPE NUM earliest earliestv ∧ MAP toks_fast lines = ss ⇒ app (p : 'ffi ffi_proj) ^(fetch_v "check_unsat''" (get_ml_prog_state())) - [fnsv; fdv; lnov; fmlv; indsv; pcv] + [fnsv; fdv; lnov; fmlv; indsv; earliestv; pcv] (STDIO fs * INSTREAM_LINES fd fdv lines fs * ARRAY fmlv fmllsv) (POSTve (λv. @@ -2433,7 +2433,7 @@ Theorem check_unsat''_spec: INSTREAM_LINES fd fdv lines' (forwardFD fs fd k) * ARRAY fmlv' fmllsv' * &( - parse_and_run fns ss fmlls inds pc = + parse_and_run fns ss fmlls inds earliest pc = SOME (MAP toks_fast lines',res) ∧ PAIR_TYPE NUM ( PAIR_TYPE (LIST_TYPE (SUM_TYPE STRING_TYPE INT)) ( @@ -2441,13 +2441,14 @@ Theorem check_unsat''_spec: PAIR_TYPE (λl v. LIST_REL (OPTION_TYPE bconstraint_TYPE) l fmllsv' ∧ v = fmlv') - (PAIR_TYPE (LIST_TYPE NUM) NPBC_CHECK_PROOF_CONF_TYPE)))) (lno',res) v)) + (PAIR_TYPE (LIST_TYPE NUM) + (NPBC_CHECK_PROOF_CONF_TYPE))))) (lno',res) v)) (λe. SEP_EXISTS k lines' fmlv' fmllsv'. ARRAY fmlv' fmllsv' * STDIO (forwardFD fs fd k) * INSTREAM_LINES fd fdv lines' (forwardFD fs fd k) * &(Fail_exn e ∧ - parse_and_run fns ss fmlls inds pc = NONE))) + parse_and_run fns ss fmlls inds earliest pc = NONE))) Proof ho_match_mp_tac (fetch "-" "parse_and_run_ind")>> rw[]>> @@ -2547,7 +2548,7 @@ Proof qexists_tac`x'`>> simp[]>> qexists_tac`k+x`>> -xsimpl)>> + xsimpl)>> simp[Once parse_and_run_def]>> qexists_tac`k+x`>>qexists_tac`x'`>>xsimpl>> qmatch_goalsub_abbrev_tac`ARRAY A B`>> @@ -2625,6 +2626,38 @@ QED val _ = translate rev_enum_full_def; +val res = translate npbc_listTheory.update_earliest_def; + +Definition fold_update_earliest_enum_def: + (fold_update_earliest_enum k [] acc = acc) ∧ + (fold_update_earliest_enum k (x::xs) acc = + fold_update_earliest_enum (k+1) + xs (update_earliest acc k (FST x))) +End + +Definition fold_update_earliest_enum_full_def: + fold_update_earliest_enum_full k fml = + fold_update_earliest_enum k fml LN +End + +Theorem fold_update_earliest_enum_FOLDL: + ∀xs k acc. + fold_update_earliest_enum k xs acc = + (FOLDL (λacc (i,v). update_earliest acc i (FST v)) acc (enumerate k xs)) +Proof + Induct>>rw[fold_update_earliest_enum_def,miscTheory.enumerate_def] +QED + +Theorem fold_update_earliest_enum_full_FOLDL: + fold_update_earliest_enum_full k xs = + (FOLDL (λacc (i,v). update_earliest acc i (FST v)) LN (enumerate k xs)) +Proof + rw[fold_update_earliest_enum_full_def,fold_update_earliest_enum_FOLDL] +QED + +val res = translate fold_update_earliest_enum_def; +val res = translate fold_update_earliest_enum_full_def; + val res = translate parse_unsat_def; val parse_unsat_side = Q.prove( @@ -2903,9 +2936,10 @@ val check_unsat' = process_topdecs ` val arr = Array.array (2*id) None val arr = fill_arr arr 1 fml val inds = rev_enum_full 1 fml + val earliest = fold_update_earliest_enum_full 1 fml val pc = init_conf id True obj in - (case check_unsat'' fns fd lno arr inds pc of + (case check_unsat'' fns fd lno arr inds earliest pc of (lno', (s, (fns',( (fml', (inds', pc')))))) => conv_boutput_hconcl @@ -2918,20 +2952,22 @@ val check_unsat' = process_topdecs ` end` |> append_prog; Theorem parse_and_run_check_csteps_list: - ∀fns ss fml inds pc rest s fns' res. - parse_and_run fns ss fml inds pc = SOME (rest, s, fns', res) ⇒ - ∃csteps. - check_csteps_list csteps fml inds pc = SOME res + ∀fns ss fml inds earliest pc rest s fns' fml' inds' pc'. + parse_and_run fns ss fml inds earliest pc = SOME (rest, s, fns', (fml', inds', pc')) ⇒ + ∃csteps earliest'. + check_csteps_list csteps fml inds earliest pc = SOME (fml', inds', earliest', pc') Proof ho_match_mp_tac parse_and_run_ind>> rw[]>> pop_assum mp_tac>> simp[Once parse_and_run_def]>> every_case_tac>>fs[] - >- - (rw[]>>qexists_tac`[]`>>EVAL_TAC>>metis_tac[])>> + >- ( + rw[]>> + qexists_tac`[]`>> + simp[npbc_listTheory.check_csteps_list_def])>> rw[]>> - first_x_assum drule>> + first_x_assum drule_all>> rw[]>> qexists_tac`y::csteps`>> simp[npbc_listTheory.check_csteps_list_def] @@ -3012,8 +3048,9 @@ Proof `BOOL T (Conv (SOME (TypeStamp "True" 0)) [])` by EVAL_TAC>> xlet_autop>> + qmatch_asmsub_abbrev_tac`SPTREE_SPT_TYPE NUM earliest earliestv`>> Cases_on` - parse_and_run fns (MAP toks_fast lines) fmlls inds + parse_and_run fns (MAP toks_fast lines) fmlls inds earliest (init_conf (LENGTH fml + 1) T obj)` >- ( (* fail to parse and run *) @@ -3068,7 +3105,7 @@ Proof ARRAY fmlv' fmllsv' * &( parse_and_run fns (MAP toks_fast lines) - fmlls inds (init_conf (LENGTH fml + 1) T obj) = + fmlls inds earliest (init_conf (LENGTH fml + 1) T obj) = SOME (MAP toks_fast lines',res) ∧ PAIR_TYPE NUM ( PAIR_TYPE (LIST_TYPE (SUM_TYPE STRING_TYPE INT)) ( @@ -3154,13 +3191,13 @@ Proof match_mp_tac (GEN_ALL npbc_listTheory.check_csteps_list_concl)>> first_x_assum (irule_at Any)>> unabbrev_all_tac>> - gs[rev_enum_full_rev_enumerate]>> + gs[rev_enum_full_rev_enumerate, fold_update_earliest_enum_full_FOLDL]>> metis_tac[])>> simp[get_bound_def]>> match_mp_tac (GEN_ALL npbc_listTheory.check_csteps_list_output)>> first_x_assum (irule_at Any)>> unabbrev_all_tac>> - gs[rev_enum_full_rev_enumerate]>> + gs[rev_enum_full_rev_enumerate, fold_update_earliest_enum_full_FOLDL]>> metis_tac[])>> metis_tac[STDIO_INSTREAM_LINES_refl_gc])>> xsimpl From a303f7b96533baccfbf2d9a67ac603058d8f6e4c Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Wed, 24 Jan 2024 22:22:49 +0800 Subject: [PATCH 09/38] anothre optimization on get_earliest --- examples/pseudo_bool/array/npbc_listScript.sml | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 0a0b5a9227..517453ceb4 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -704,7 +704,9 @@ Definition check_red_list_fast_def: End Definition get_earliest_def: - (get_earliest earliest (INR _) = SOME (0:num)) ∧ + (get_earliest earliest (INR v) = + if length v = 0 then NONE + else SOME (0:num)) ∧ (get_earliest earliest (INL (n,_)) = sptree$lookup n earliest) End From 1eeed16f3f68d94da53896065ec4112182c24154 Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Fri, 26 Jan 2024 15:56:51 +0800 Subject: [PATCH 10/38] replace earliest with vimap --- .../pseudo_bool/array/npbc_listScript.sml | 398 ++++++++++++------ 1 file changed, 261 insertions(+), 137 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 517453ceb4..426ffe56e4 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -690,7 +690,7 @@ End (* inds is just passed through here *) Definition check_red_list_fast_def: - check_red_list_fast b fml inds id c pf cid = + check_red_list_fast b fml inds id c pf cid vimap = let nc = not c in let fml_not_c = update_resize fml NONE (SOME (nc,b)) id in case check_lsteps_list pf b fml_not_c id (id+1) of @@ -698,27 +698,49 @@ Definition check_red_list_fast_def: | SOME (fml', id') => if check_contradiction_fml_list b fml' cid then let rfml = rollback fml' id id' in - SOME (rfml,inds,id') + SOME (rfml,inds,vimap,id') else NONE End +(* Definition get_earliest_def: (get_earliest earliest (INR v) = if length v = 0 then NONE else SOME (0:num)) ∧ (get_earliest earliest (INL (n,_)) = sptree$lookup n earliest) +End *) + +(* A reverse mapping of vars -> indices *) +Definition get_indices_def: + get_indices b tcb fml inds s vimap = + case s of + INR v => + if length v = 0 then ([],[]) + else (reindex (b ∨ tcb) fml inds) + | INL (n,_) => + case sptree$lookup n vimap of + NONE => ([],[]) + | SOME inds => (reindex (b ∨ tcb) fml inds) +End + +Definition set_indices_def: + set_indices inds s vimap rinds = + case s of + INR v => + if length v = 0 then (inds,vimap) + else (rinds,vimap) + | INL (n,_) => + (inds, sptree$insert n rinds vimap) End Definition check_red_list_def: - check_red_list ord obj b tcb fml inds id c s pfs idopt earliest = + check_red_list ord obj b tcb fml inds id c s pfs idopt vimap = let s = mk_subst s in case red_fast s idopt pfs of NONE => ( - let mini = get_earliest earliest s in - let (rinds,fmlls,rest) = - reindex_partial (b ∨ tcb) fml mini inds in + let (rinds,fmlls) = get_indices b tcb fml inds s vimap in let nc = not c in let fml_not_c = update_resize fml NONE (SOME (nc,b)) id in let rsubs = red_subgoals ord (subst_fun s) c obj in @@ -732,19 +754,22 @@ Definition check_red_list_def: let rfml = rollback fml' id id' in if do_red_check idopt b tcb fml' s rfml rinds fmlls nc pfs rsubs then - SOME (rfml,rinds ++ rest,id') + let (inds',vimap') = set_indices inds s vimap rinds in + SOME (rfml,inds',vimap',id') else NONE)) | SOME (pf,cid) => - check_red_list_fast b fml inds id c pf cid + check_red_list_fast b fml inds id c pf cid vimap End +(* Definition min_opt_def: min_opt i j = case i of NONE => j | SOME ii => MIN ii j -End +End *) (* v is the new index of the constraint (last arg) *) +(* Definition update_earliest_def: (update_earliest earliest v [] = earliest) ∧ (update_earliest earliest v ((i,n)::ns) = @@ -753,34 +778,49 @@ Definition update_earliest_def: v ns) End +*) + +Definition opt_cons_def: + (opt_cons v NONE = [v]) ∧ + (opt_cons v (SOME ls) = v::ls) +End + +Definition update_vimap_def: + (update_vimap vimap v [] = vimap) ∧ + (update_vimap vimap v ((i,n)::ns) = + update_vimap + (insert n (opt_cons v (lookup n vimap)) vimap) + v + ns) +End Definition opt_update_inds_def[simp]: - (opt_update_inds fml NONE id inds earliest = - (fml,inds,earliest,id)) ∧ - (opt_update_inds fml (SOME cc) id inds earliest = + (opt_update_inds fml NONE id inds vimap = + (fml,inds,vimap,id)) ∧ + (opt_update_inds fml (SOME cc) id inds vimap = (update_resize fml NONE (SOME cc) id, sorted_insert id inds, - update_earliest earliest id (FST (FST cc)), + update_vimap vimap id (FST (FST cc)), id+1)) End Definition check_sstep_list_def: (check_sstep_list (sstep:sstep) ord obj tcb (fml: (npbc # bool) option list) (inds:num list) (id:num) - earliest = + vimap = case sstep of | Lstep lstep => (case check_lstep_list lstep F fml 0 id of NONE => NONE | SOME (rfml,c,id') => - SOME (opt_update_inds rfml c id' inds earliest)) + SOME (opt_update_inds rfml c id' inds vimap)) | Red c s pfs idopt => case check_red_list ord obj F tcb fml inds id c s pfs - idopt earliest of - SOME (rfml,rinds,id') => + idopt vimap of + SOME (rfml,rinds,vimap',id') => SOME ( update_resize rfml NONE (SOME (c,tcb)) id', sorted_insert id' rinds, - update_earliest earliest id' (FST c), + update_vimap vimap' id' (FST c), id'+1) | NONE => NONE) End @@ -1314,6 +1354,7 @@ Proof simp[lookup_core_only_list_def,any_el_ALT,EL_LUPDATE] QED +(* Definition earliest_rel_def: earliest_rel fmlls earliest ⇔ ∀x pos. @@ -1321,7 +1362,18 @@ Definition earliest_rel_def: case EL pos fmlls of NONE => T | SOME c => ¬MEM x (MAP SND (FST (FST c))) End +*) + +Definition vimap_rel_def: + vimap_rel fmlls vimap ⇔ + ∀i c x. + i < LENGTH fmlls ∧ + EL i fmlls = SOME c ∧ + MEM x (MAP SND (FST (FST c))) ⇒ + ∃ls. sptree$lookup x vimap = SOME ls ∧ MEM i ls +End +(* (* If we already proved fml_rel, then we get earliest_rel for free *) Theorem fml_rel_fml_rel_earliest_rel: @@ -1338,21 +1390,55 @@ Proof gvs[min_opt_def]>> every_case_tac>>fs[] QED +*) + +Theorem fml_rel_fml_rel_vimap_rel: + fml_rel fml fmlls ∧ + vimap_rel fmlls vimap ∧ + fml_rel fml fmlls' ⇒ + vimap_rel fmlls' vimap +Proof + rw[fml_rel_def,vimap_rel_def]>> + first_x_assum(qspec_then `i` mp_tac)>> + last_x_assum(qspec_then `i` mp_tac)>> + rw[any_el_ALT]>>gvs[]>> + metis_tac[] +QED + +Theorem ind_rel_get_indices_set_indices: + get_indices b tcb fmlls inds s vimap = (rinds,_) ∧ + set_indices inds s vimap rinds = (inds',vimap') ∧ + ind_rel fmlls inds ⇒ + ind_rel fmlls inds' +Proof + rw[get_indices_def,set_indices_def] >> + gvs[AllCaseEqs()]>> + metis_tac[ind_rel_reindex] +QED + +Theorem vimap_rel_get_indices_set_indices: + get_indices b tcb fmlls inds s vimap = (rinds,_) ∧ + set_indices inds s vimap rinds = (inds',vimap') ∧ + vimap_rel fmlls vimap ⇒ + vimap_rel fmlls vimap' +Proof + rw[get_indices_def,set_indices_def] >> + gvs[AllCaseEqs()]>> + cheat +QED Theorem fml_rel_check_red_list: fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ - SORTED $>= inds ∧ - earliest_rel fmlls earliest ∧ + vimap_rel fmlls vimap ∧ (∀n. n ≥ id ⇒ any_el n fmlls NONE = NONE) ∧ check_red_list ord obj b tcb fmlls inds id c s pfs - idopt earliest = - SOME (fmlls', inds', id') ⇒ + idopt vimap = + SOME (fmlls', inds', vimap', id') ⇒ check_red ord obj b tcb fml id c s pfs idopt = SOME id' ∧ fml_rel fml fmlls' ∧ ind_rel fmlls' inds' ∧ - SORTED $>= inds' ∧ - earliest_rel fmlls' earliest ∧ + vimap_rel fmlls' vimap' ∧ (∀n. n ≥ id' ⇒ any_el n fmlls' NONE = NONE) ∧ id ≤ id' Proof @@ -1360,6 +1446,7 @@ Proof fs[check_red_list_def]>> gvs[AllCaseEqs()] >- ( + pairarg_tac>>fs[]>> pairarg_tac>>fs[]>> every_case_tac>>gvs[]>> simp[check_red_def]>> @@ -1389,7 +1476,8 @@ Proof disch_then (qspec_then`mk_core_fml (b ∨ tcb) fml` mp_tac)>> impl_tac >- ( simp[range_mk_core_fml]>> - match_mp_tac (GEN_ALL SND_reindex_partial_characterize)>> + gvs[get_indices_def] >> every_case_tac>> gvs[]>> + match_mp_tac (GEN_ALL SND_reindex_characterize)>> metis_tac[])>> match_mp_tac split_goals_same_goals>> simp[EXTENSION,FORALL_PROD]>> @@ -1432,6 +1520,8 @@ Proof metis_tac[])*))>> drule subst_indexes_MEM>> rw[MEM_toAList,lookup_map_opt]>> + cheat + (* drule_all FST_reindex_partial_characterize>> TOP_CASE_TAC>>rw[]>>gvs[]>> fs[rollback_def,lookup_core_only_list_list_delete_list,MEM_MAP,MEM_COUNT_LIST,MEM_FILTER]>> @@ -1446,7 +1536,7 @@ Proof drule (GSYM fml_rel_lookup_core_only)>> strip_tac>>fs[]>> gvs[lookup_core_only_list_def,AllCaseEqs()])>> - fs[])>> + fs[]*))>> match_mp_tac (GEN_ALL fml_rel_check_contradiction_fml)>> metis_tac[])>> CONJ_ASM1_TAC>- ( @@ -1454,11 +1544,10 @@ Proof CONJ_TAC >- ( match_mp_tac ind_rel_rollback_2>> simp[] >> - metis_tac[ind_rel_reindex_partial])>> - CONJ_TAC >- - metis_tac[ind_rel_reindex_partial]>> - CONJ_TAC >- - metis_tac[fml_rel_fml_rel_earliest_rel]>> + metis_tac[ind_rel_get_indices_set_indices])>> + CONJ_TAC >- ( + drule_all vimap_rel_get_indices_set_indices>> + metis_tac[fml_rel_fml_rel_vimap_rel])>> simp[rollback_def,any_el_list_delete_list,MEM_MAP,MEM_COUNT_LIST]>> rw[])>> gvs[check_red_list_fast_def,AllCaseEqs(),check_red_def,red_fast_def,extract_clauses_def,check_subproofs_def,insert_fml_def,check_lstep_list_def] @@ -1478,8 +1567,8 @@ Proof match_mp_tac ind_rel_rollback_2>> simp[any_el_update_resize])>> CONJ_TAC - >- (* earliest_rel *) - metis_tac[fml_rel_fml_rel_earliest_rel] + >- (* vimap_rel *) + metis_tac[fml_rel_fml_rel_vimap_rel] >- simp[rollback_def,any_el_list_delete_list,MEM_MAP,MEM_COUNT_LIST,any_el_update_resize])>> `fml_rel (insert id ((not c,b)) fml) @@ -1503,15 +1592,15 @@ Proof CONJ_TAC >- ( match_mp_tac ind_rel_rollback_2>> simp[])>> - CONJ_TAC >- (* earliest_rel *) - metis_tac[fml_rel_fml_rel_earliest_rel]>> + CONJ_TAC >- (* vimap_rel *) + metis_tac[fml_rel_fml_rel_vimap_rel]>> simp[rollback_def,any_el_list_delete_list,MEM_MAP,MEM_COUNT_LIST,any_el_update_resize]>> rw[] QED Theorem opt_update_inds_opt_update: - opt_update_inds fml c id inds earliest = - (fml',inds',earliest',id') ⇒ + opt_update_inds fml c id inds vimap = + (fml',inds',vimap',id') ⇒ opt_update fml c id = (fml',id') Proof Cases_on`c`>>rw[] @@ -1536,8 +1625,8 @@ QED Theorem opt_update_inds_SORTED: SORTED $>= inds ∧ - opt_update_inds fml c id inds earliest = - (fml',inds',earliest',id') ⇒ + opt_update_inds fml c id inds vimap = + (fml',inds',vimap',id') ⇒ SORTED $>= inds' Proof Cases_on`c`>>rw[]>>fs[]>> @@ -1546,14 +1635,15 @@ QED Theorem opt_update_inds_ind_rel: ind_rel fml inds ∧ - opt_update_inds fml c id inds earliest = - (fml',inds',earliest',id') ⇒ + opt_update_inds fml c id inds vimap = + (fml',inds',vimap',id') ⇒ ind_rel fml' inds' Proof Cases_on`c`>>rw[]>>fs[]>> simp[ind_rel_update_resize_sorted_insert] QED +(* Theorem earliest_rel_check_lstep_list: earliest_rel fmlls earliest ∧ (∀n. n ≥ id ⇒ any_el n fmlls NONE = NONE) ∧ @@ -1671,22 +1761,60 @@ Proof Cases_on`c`>>rw[]>>fs[]>> metis_tac[earliest_rel_update_resize_update_earliest,FST,PAIR] QED +*) + +Theorem vimap_rel_check_lstep_list: + vimap_rel fmlls vimap ∧ + (∀n. n ≥ id ⇒ any_el n fmlls NONE = NONE) ∧ + check_lstep_list lstep b fmlls mindel id = + SOME (fmlls',x,y) ⇒ + vimap_rel fmlls' vimap +Proof + rw[]>> + fs[vimap_rel_def]>> + rw[]>> + drule (CONJUNCT1 check_lstep_list_id_del)>> + drule (CONJUNCT1 check_lstep_list_id_upper)>> + disch_then drule>>rw[]>> + `i < id` by ( + gvs[any_el_ALT,AND_IMP_INTRO]>> + first_x_assum (drule_at Any)>> + simp[])>> + first_x_assum drule>> + simp[any_el_ALT] +QED + +Theorem vimap_rel_update_resize_update_vimap: + vimap_rel fml vimap ⇒ + vimap_rel (update_resize fml NONE (SOME (v,b)) n) + (update_vimap vimap n (FST v)) +Proof + cheat +QED + +Theorem opt_update_inds_vimap_rel: + vimap_rel fml vimap ∧ + opt_update_inds fml c id inds vimap = + (fml',inds',vimap',id') ⇒ + vimap_rel fml' vimap' +Proof + Cases_on`c`>>rw[]>>fs[]>> + metis_tac[vimap_rel_update_resize_update_vimap,FST,PAIR] +QED Theorem fml_rel_check_sstep_list: ∀sstep ord obj fmlls inds id fmlls' id' inds' fml. fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ - SORTED $>= inds ∧ - earliest_rel fmlls earliest ∧ + vimap_rel fmlls vimap ∧ (∀n. n ≥ id ⇒ any_el n fmlls NONE = NONE) ∧ - check_sstep_list sstep ord obj tcb fmlls inds id earliest = - SOME (fmlls',inds',earliest',id') ⇒ + check_sstep_list sstep ord obj tcb fmlls inds id vimap = + SOME (fmlls',inds',vimap',id') ⇒ ∃fml'. check_sstep sstep ord obj tcb fml id = SOME(fml',id') ∧ fml_rel fml' fmlls' ∧ ind_rel fmlls' inds' ∧ - SORTED $>= inds' ∧ - earliest_rel fmlls' earliest' ∧ + vimap_rel fmlls' vimap' ∧ (∀n. n ≥ id' ⇒ any_el n fmlls' NONE = NONE) ∧ id ≤ id' Proof @@ -1701,9 +1829,7 @@ Proof CONJ_TAC >- metis_tac[opt_update_inds_ind_rel,ind_rel_check_lstep_list]>> CONJ_TAC >- - metis_tac[opt_update_inds_SORTED]>> - CONJ_TAC >- - metis_tac[opt_update_inds_earliest_rel,earliest_rel_check_lstep_list]>> + metis_tac[opt_update_inds_vimap_rel,vimap_rel_check_lstep_list]>> drule (CONJUNCT1 check_lstep_list_id_upper)>> drule opt_update_id_upper>> drule opt_update_id>> @@ -1719,9 +1845,7 @@ Proof >- metis_tac[ind_rel_update_resize_sorted_insert] >- - metis_tac[SORTED_sorted_insert] - >- - metis_tac[PAIR,FST,SND,earliest_rel_update_resize_update_earliest]>> + metis_tac[PAIR,FST,SND,vimap_rel_update_resize_update_vimap]>> simp[any_el_update_resize]) QED @@ -1842,7 +1966,7 @@ Definition check_change_obj_list_def: End Definition check_cstep_list_def: - check_cstep_list cstep fml inds earliest pc = + check_cstep_list cstep fml inds vimap pc = case cstep of Dom c s pfs idopt => (case pc.ord of @@ -1867,14 +1991,14 @@ Definition check_cstep_list_def: SOME( update_resize rfml NONE (SOME (c,pc.tcb)) id', sorted_insert id' rinds, - update_earliest earliest id' (FST c), + update_vimap vimap id' (FST c), pc with id := id'+1) else NONE))) | Sstep sstep => (case check_sstep_list sstep pc.ord pc.obj pc.tcb - fml inds pc.id earliest of - SOME(fml',inds',earliest',id') => - SOME(fml',inds', earliest',pc with id := id') + fml inds pc.id vimap of + SOME(fml',inds',vimap',id') => + SOME(fml',inds', vimap',pc with id := id') | NONE => NONE) | CheckedDelete n s pfs idopt => ( if check_tcb_idopt pc.tcb idopt then @@ -1883,9 +2007,9 @@ Definition check_cstep_list_def: | SOME c => (let nfml = delete_list n fml in case check_red_list pc.ord pc.obj T pc.tcb - nfml inds pc.id c s pfs idopt earliest of - SOME (ncf',inds',id') => - SOME (ncf', inds', earliest, pc with <| id := id' |>) + nfml inds pc.id c s pfs idopt vimap of + SOME (ncf',inds',vimap',id') => + SOME (ncf', inds', vimap', pc with <| id := id' |>) | NONE => NONE) ) else NONE) | UncheckedDelete ls => ( @@ -1893,25 +2017,25 @@ Definition check_cstep_list_def: if ¬pc.tcb ∧ pc.ord = NONE then SOME (list_delete_list ls fml, inds, - earliest, pc with chk := F) + vimap, pc with chk := F) else case all_core_list fml inds [] of NONE => NONE | SOME inds' => SOME (list_delete_list ls fml, inds', - earliest, pc with chk := F)) + vimap, pc with chk := F)) | Transfer ls => (case core_from_inds fml ls of NONE => NONE | SOME fml' => - SOME (fml', inds, earliest, pc)) + SOME (fml', inds, vimap, pc)) | StrengthenToCore b => (let inds' = FST (reindex F fml inds) in let pc' = pc with tcb := b in if b then (case core_from_inds fml inds' of NONE => NONE - | SOME fml' => SOME (fml',inds', earliest, pc')) + | SOME fml' => SOME (fml',inds', vimap, pc')) else - SOME (fml,inds',earliest,pc')) + SOME (fml,inds',vimap,pc')) | LoadOrder nn xs => (let inds' = FST (reindex F fml inds) in case ALOOKUP pc.orders nn of NONE => NONE @@ -1919,12 +2043,12 @@ Definition check_cstep_list_def: if LENGTH xs = LENGTH (FST (SND ord')) then case core_from_inds fml inds' of NONE => NONE | SOME fml' => - SOME (fml',inds',earliest,pc with ord := SOME (ord',xs)) + SOME (fml',inds',vimap,pc with ord := SOME (ord',xs)) else NONE) | UnloadOrder => (case pc.ord of NONE => NONE | SOME spo => - SOME (fml,inds, earliest, pc with ord := NONE)) + SOME (fml,inds, vimap, pc with ord := NONE)) | StoreOrder nn spo ws pfsr pfst => if check_good_ord spo ∧ check_ws spo ws then @@ -1932,7 +2056,7 @@ Definition check_cstep_list_def: | SOME id => if check_reflexivity spo pfsr id then SOME (fml, inds, - earliest, pc with orders := (nn,spo)::pc.orders) + vimap, pc with orders := (nn,spo)::pc.orders) else NONE else NONE @@ -1948,13 +2072,13 @@ Definition check_cstep_list_def: SOME ( update_resize fml NONE (SOME (c,T)) pc.id, sorted_insert pc.id inds, - update_earliest earliest pc.id (FST c), + update_vimap vimap pc.id (FST c), pc with <| id := pc.id+1; bound := bound'; dbound := dbound' |>) else - SOME (fml, inds, earliest, + SOME (fml, inds, vimap, pc with <| bound := bound'; dbound := dbound' |>)) @@ -1963,11 +2087,11 @@ Definition check_cstep_list_def: NONE => NONE | SOME (fml',fc',id') => SOME ( - fml', inds, earliest, + fml', inds, vimap, pc with <| id:=id'; obj:=SOME fc' |>)) | CheckObj fc' => if check_eq_obj pc.obj fc' - then SOME (fml, inds, earliest, pc) + then SOME (fml, inds, vimap, pc) else NONE End @@ -2068,18 +2192,20 @@ Proof simp[rollback_def,any_el_list_delete_list,MEM_MAP,MEM_COUNT_LIST] QED -Theorem earliest_rel_list_delete_list: +Theorem vimap_rel_list_delete_list: ∀l fmlls. - earliest_rel fmlls earliest ==> - earliest_rel (list_delete_list l fmlls) earliest + vimap_rel fmlls vimap ==> + vimap_rel (list_delete_list l fmlls) vimap Proof Induct \\ gvs [list_delete_list_def] \\ rw [] \\ last_x_assum irule - \\ gvs [earliest_rel_def] + \\ gvs [vimap_rel_def] \\ ‘LENGTH (delete_list h fmlls) = LENGTH fmlls’ by rw [delete_list_def] \\ gvs [] \\ rw [] - \\ last_x_assum drule - \\ rw [delete_list_def] \\ rw [EL_LUPDATE] + \\ last_x_assum irule \\ simp[] + \\ last_x_assum (irule_at Any) + \\ pop_assum mp_tac + \\ rw [delete_list_def,EL_LUPDATE] QED Theorem all_core_list_SORTED: @@ -2095,17 +2221,15 @@ QED Theorem fml_rel_check_cstep_list: fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ - SORTED $>= inds ∧ - earliest_rel fmlls earliest ∧ + vimap_rel fmlls vimap ∧ (∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE) ∧ - check_cstep_list cstep fmlls inds earliest pc = - SOME (fmlls',inds',earliest',pc') ⇒ + check_cstep_list cstep fmlls inds vimap pc = + SOME (fmlls',inds',vimap',pc') ⇒ ∃fml'. check_cstep cstep fml pc = SOME (fml', pc') ∧ fml_rel fml' fmlls' ∧ ind_rel fmlls' inds' ∧ - SORTED $>= inds' ∧ - earliest_rel fmlls' earliest' ∧ + vimap_rel fmlls' vimap' ∧ (∀n. n ≥ pc'.id ⇒ any_el n fmlls' NONE = NONE) ∧ pc.id ≤ pc'.id Proof @@ -2151,11 +2275,9 @@ Proof match_mp_tac ind_rel_rollback_2>> fs[]>> metis_tac[ind_rel_reindex])>> - CONJ_TAC >- - metis_tac[SORTED_sorted_insert,SORTED_reindex]>> CONJ_TAC >- ( - match_mp_tac earliest_rel_update_resize_update_earliest>> - match_mp_tac fml_rel_fml_rel_earliest_rel>>fs[]>> + match_mp_tac vimap_rel_update_resize_update_vimap>> + match_mp_tac fml_rel_fml_rel_vimap_rel>>fs[]>> match_mp_tac fml_rel_rollback>>rw[]>>fs[])>> simp[rollback_def,any_el_list_delete_list,MEM_MAP,MEM_COUNT_LIST]) >- ( (* Sstep *) @@ -2181,7 +2303,7 @@ Proof disch_then(qspec_then`[n]` mp_tac)>> simp[list_delete_list_def])>> CONJ_TAC >- ( - drule earliest_rel_list_delete_list>> + drule vimap_rel_list_delete_list>> disch_then(qspec_then`[n]` mp_tac)>> simp[list_delete_list_def])>> metis_tac[any_el_list_delete_list,list_delete_list_def]) @@ -2193,7 +2315,7 @@ Proof CONJ_TAC >- metis_tac[ind_rel_list_delete_list]>> CONJ_TAC >- - metis_tac[earliest_rel_list_delete_list]>> + metis_tac[vimap_rel_list_delete_list]>> simp[any_el_list_delete_list]) >- ( drule_all fml_rel_all_core>>strip_tac>> @@ -2203,9 +2325,7 @@ Proof CONJ_TAC >- metis_tac[ind_rel_list_delete_list]>> CONJ_TAC >- - metis_tac[all_core_list_SORTED]>> - CONJ_TAC >- - metis_tac[earliest_rel_list_delete_list]>> + metis_tac[vimap_rel_list_delete_list]>> simp[any_el_list_delete_list])) >- ( (* Transfer *) gvs[check_cstep_list_def,AllCaseEqs(),check_cstep_def]>> @@ -2215,13 +2335,12 @@ Proof fs[ind_rel_def]>> rw[]>> (* TODO should be an easy induction *) - `earliest_rel fmlls' earliest` by cheat>> + `vimap_rel fmlls' vimap` by cheat>> metis_tac[IS_SOME_EXISTS,option_CLAUSES]) >- ( (* StrengthenToCore *) gvs[check_cstep_list_def,AllCaseEqs(),check_cstep_def]>> Cases_on`reindex F fmlls inds`>> - drule_all ind_rel_reindex>> - drule_all SORTED_reindex + drule_all ind_rel_reindex >- ( drule any_el_core_from_inds>> rw[] @@ -2238,19 +2357,26 @@ Proof fs[ind_rel_def]>> rw[]>> metis_tac[IS_SOME_EXISTS,option_CLAUSES]) - >- - cheat) + >- ( + fs[vimap_rel_def]>>rw[]>> + first_x_assum match_mp_tac>> + first_x_assum(qspec_then`i` mp_tac)>> + rw[any_el_ALT]>>gvs[]>> + fs[MEM_MAP] + >- ( + pairarg_tac>>fs[]>> + metis_tac[FST,PAIR,SND])>> + metis_tac[FST,PAIR,SND]) + ) >- fs[]) >- ( (* LoadOrder *) gvs[check_cstep_list_def,AllCaseEqs(),check_cstep_def]>> Cases_on`reindex F fmlls inds`>> drule_all ind_rel_reindex>> - drule_all SORTED_reindex>> drule any_el_core_from_inds>> strip_tac>>fs[]>> strip_tac>>fs[]>> - strip_tac>>fs[]>> CONJ_TAC >- ( simp[fml_rel_def,lookup_map]>> fs[ind_rel_def]>>rw[] @@ -2282,8 +2408,7 @@ Proof rw[]>>gvs[] >- metis_tac[fml_rel_update_resize] >- metis_tac[ind_rel_update_resize_sorted_insert] - >- metis_tac[SORTED_sorted_insert] - >- metis_tac[earliest_rel_update_resize_update_earliest]>> + >- metis_tac[vimap_rel_update_resize_update_vimap]>> simp[any_el_update_resize]) >- ( (* ChangeObj *) fs[check_cstep_def,check_cstep_list_def]>> @@ -2316,7 +2441,7 @@ Proof simp[]>> metis_tac[ind_rel_reindex])>> CONJ_TAC >- - metis_tac[fml_rel_fml_rel_earliest_rel]>> + metis_tac[fml_rel_fml_rel_vimap_rel]>> simp[any_el_rollback]) >- ( (* CheckObj *) fs[check_cstep_def,check_cstep_list_def] @@ -2324,30 +2449,28 @@ Proof QED Definition check_csteps_list_def: - (check_csteps_list [] fml inds earliest pc = - SOME (fml, inds, earliest, pc)) ∧ - (check_csteps_list (c::cs) fml inds earliest pc = - case check_cstep_list c fml inds earliest pc of + (check_csteps_list [] fml inds vimap pc = + SOME (fml, inds, vimap, pc)) ∧ + (check_csteps_list (c::cs) fml inds vimap pc = + case check_cstep_list c fml inds vimap pc of NONE => NONE - | SOME(fml', inds', earliest', pc') => - check_csteps_list cs fml' inds' earliest' pc') + | SOME(fml', inds', vimap', pc') => + check_csteps_list cs fml' inds' vimap' pc') End Theorem fml_rel_check_csteps_list: - ∀csteps fml fmlls inds earliest pc fmlls' inds' earliest' pc'. + ∀csteps fml fmlls inds vimap pc fmlls' inds' vimap' pc'. fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ - SORTED $>= inds ∧ - earliest_rel fmlls earliest ∧ + vimap_rel fmlls vimap ∧ (∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE) ∧ - check_csteps_list csteps fmlls inds earliest pc = - SOME (fmlls', inds', earliest', pc') ⇒ + check_csteps_list csteps fmlls inds vimap pc = + SOME (fmlls', inds', vimap', pc') ⇒ ∃fml'. check_csteps csteps fml pc = SOME (fml', pc') ∧ fml_rel fml' fmlls' ∧ ind_rel fmlls' inds' ∧ - SORTED $>= inds' ∧ - earliest_rel fmlls' earliest' ∧ + vimap_rel fmlls' vimap' ∧ (∀n. n ≥ pc'.id ⇒ any_el n fmlls' NONE = NONE) ∧ pc.id ≤ pc'.id Proof @@ -2544,27 +2667,28 @@ Proof gs[EL_REPLICATE] QED -Theorem earliest_rel_FOLDL_update_resize_aux: +Theorem vimap_rel_FOLDL_update_resize_aux: ∀xs ls t. - earliest_rel ls t ⇒ - earliest_rel + vimap_rel ls t ⇒ + vimap_rel (FOLDL (λacc (i,v). update_resize acc NONE (SOME (v,b)) i) ls xs) - (FOLDL (λacc (i,v). update_earliest acc i (FST v)) t xs) + (FOLDL (λacc (i,v). update_vimap acc i (FST v)) t xs) Proof Induct>>rw[]>> first_x_assum match_mp_tac>> pairarg_tac>>gvs[]>> - match_mp_tac earliest_rel_update_resize_update_earliest>> + match_mp_tac vimap_rel_update_resize_update_vimap>> fs[] QED -Theorem earliest_rel_FOLDL_update_resize: - earliest_rel +Theorem vimap_rel_FOLDL_update_resize: + vimap_rel (FOLDL (λacc (i,v). update_resize acc NONE (SOME (v,b)) i) (REPLICATE n NONE) (enumerate k fml)) - (FOLDL (λacc (i,v). update_earliest acc i (FST v)) LN (enumerate k fml)) + (FOLDL (λacc (i,v). update_vimap acc i (FST v)) LN (enumerate k fml)) Proof - match_mp_tac earliest_rel_FOLDL_update_resize_aux>> - simp[earliest_rel_def,min_opt_def,EL_REPLICATE] + match_mp_tac vimap_rel_FOLDL_update_resize_aux>> + rw[vimap_rel_def]>> + CCONTR_TAC>>fs[EL_REPLICATE] QED Theorem SORTED_REVERSE_enumerate: @@ -2582,16 +2706,16 @@ Theorem check_csteps_list_concl: (FOLDL (λacc (i,v). update_resize acc NONE (SOME (v,T)) i) (REPLICATE m NONE) (enumerate 1 fml)) (REVERSE (MAP FST (enumerate 1 fml))) - (FOLDL (λacc (i,v). update_earliest acc i (FST v)) LN (enumerate 1 fml)) + (FOLDL (λacc (i,v). update_vimap acc i (FST v)) LN (enumerate 1 fml)) (init_conf (LENGTH fml + 1) chk obj) = - SOME(fmlls',inds',earliest',pc') ∧ + SOME(fmlls',inds',vimap',pc') ∧ check_hconcl_list fml obj fmlls' pc'.obj pc'.bound pc'.dbound hconcl ⇒ sem_concl (set fml) obj (hconcl_concl hconcl) Proof rw[]>> qmatch_asmsub_abbrev_tac`check_csteps_list cs fmlls inds - earliest pc = _`>> + vimap pc = _`>> `fml_rel (build_fml T 1 fml) fmlls` by simp[Abbr`fmlls`,fml_rel_FOLDL_update_resize]>> `ind_rel fmlls inds` by ( @@ -2599,9 +2723,9 @@ Proof simp[ind_rel_FOLDL_update_resize])>> `SORTED $>= inds` by (unabbrev_all_tac>>fs[SORTED_REVERSE_enumerate])>> - `earliest_rel fmlls earliest` by ( + `vimap_rel fmlls vimap` by ( unabbrev_all_tac>> - simp[earliest_rel_FOLDL_update_resize])>> + simp[vimap_rel_FOLDL_update_resize])>> `∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE` by ( rw[Abbr`pc`,Abbr`fmlls`,any_el_ALT,init_conf_def]>> DEP_REWRITE_TAC [FOLDL_update_resize_lookup]>> @@ -2697,15 +2821,15 @@ Theorem check_csteps_list_output: (FOLDL (λacc (i,v). update_resize acc NONE (SOME (v,T)) i) (REPLICATE m NONE) (enumerate 1 fml)) (REVERSE (MAP FST (enumerate 1 fml))) - (FOLDL (λacc (i,v). update_earliest acc i (FST v)) LN (enumerate 1 fml)) + (FOLDL (λacc (i,v). update_vimap acc i (FST v)) LN (enumerate 1 fml)) (init_conf (LENGTH fml + 1) chk obj) = - SOME(fmlls',inds',earliest',pc') ∧ + SOME(fmlls',inds',vimap',pc') ∧ check_output_list fmlls' inds' pc'.obj pc'.bound pc'.dbound pc'.chk fmlt objt output ⇒ sem_output (set fml) obj pc'.bound (set fmlt) objt output Proof rw[]>> - qmatch_asmsub_abbrev_tac`check_csteps_list cs fmlls inds earliest pc = _`>> + qmatch_asmsub_abbrev_tac`check_csteps_list cs fmlls inds vimap pc = _`>> `fml_rel (build_fml T 1 fml) fmlls` by simp[Abbr`fmlls`,fml_rel_FOLDL_update_resize]>> `ind_rel fmlls inds` by ( @@ -2713,9 +2837,9 @@ Proof simp[ind_rel_FOLDL_update_resize])>> `SORTED $>= inds` by (unabbrev_all_tac>>fs[SORTED_REVERSE_enumerate])>> - `earliest_rel fmlls earliest` by ( + `vimap_rel fmlls vimap` by ( unabbrev_all_tac>> - simp[earliest_rel_FOLDL_update_resize])>> + simp[vimap_rel_FOLDL_update_resize])>> `∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE` by ( rw[Abbr`pc`,Abbr`fmlls`,any_el_ALT,init_conf_def]>> DEP_REWRITE_TAC [FOLDL_update_resize_lookup]>> From 75231596b94195c65c22d48f4c4c58226e6c07c7 Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Fri, 26 Jan 2024 17:42:53 +0800 Subject: [PATCH 11/38] get changes past Prog --- .../array/npbc_arrayProgScript.sml | 221 ++++++++++++------ .../array/npbc_parseProgScript.sml | 88 ++++--- 2 files changed, 196 insertions(+), 113 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_arrayProgScript.sml b/examples/pseudo_bool/array/npbc_arrayProgScript.sml index ed3bd2aa14..5253bbe59d 100644 --- a/examples/pseudo_bool/array/npbc_arrayProgScript.sml +++ b/examples/pseudo_bool/array/npbc_arrayProgScript.sml @@ -1840,7 +1840,7 @@ QED val r = translate (red_fast_def |> SIMP_RULE std_ss [vec_eq_nil_thm]); val check_red_arr_fast = process_topdecs` - fun check_red_arr_fast lno b fml inds id c pf cid = + fun check_red_arr_fast lno b fml inds id c pf cid vimap = let val nc = not_1 c val fml_not_c = Array.updateResize fml None id (Some (nc,b)) in @@ -1848,12 +1848,69 @@ val check_red_arr_fast = process_topdecs` (fml', id') => if check_contradiction_fml_arr b fml' cid then let val u = rollback_arr fml' id id' in - (fml', (inds, id')) + (fml', (inds, (vimap, id'))) end else raise Fail (format_failure lno ("did not derive contradiction from index:" ^ Int.toString cid)) end` |> append_prog; -val res = translate get_earliest_def; +val res = translate set_indices_def; + +Definition get_indices_pure_def: + get_indices_pure inds s vimap = + case s of + INR v => + if length v = 0 then INL () + else INR inds + | INL (n,_) => + case sptree$lookup n vimap of + NONE => INL () + | SOME inds => INR inds +End + +val res = translate get_indices_pure_def; + +val get_indices_arr = process_topdecs` + fun get_indices_arr bortcb fml inds s vimap = + case get_indices_pure inds s vimap of + Inl v => ([],[]) + | Inr inds => + (reindex_arr bortcb fml inds)` |> append_prog; + +Theorem get_indices_arr_spec: + BOOL (b ∨ tcb) bv ∧ + LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ + (LIST_TYPE NUM) inds indsv ∧ + subst_TYPE s sv ∧ + vimap_TYPE vimap vimapv + ⇒ + app (p : 'ffi ffi_proj) + ^(fetch_v "get_indices_arr" (get_ml_prog_state())) + [bv; fmlv; indsv; sv; vimapv] + (ARRAY fmlv fmllsv) + (POSTv v. + ARRAY fmlv fmllsv * + &( + PAIR_TYPE (LIST_TYPE NUM) + (LIST_TYPE constraint_TYPE) + (get_indices b tcb fmlls inds s vimap) v)) +Proof + rw[]>> + xcf "get_indices_arr" (get_ml_prog_state ())>> + xlet_autop>> + fs[get_indices_pure_def,get_indices_def]>> + every_case_tac>>fs[SUM_TYPE_def]>>xmatch + >- ( + rpt xlet_autop>>xcon>> + xsimpl>> + simp[PAIR_TYPE_def,LIST_TYPE_def]) + >- + (xapp>>xsimpl) + >- ( + rpt xlet_autop>>xcon>> + xsimpl>> + simp[PAIR_TYPE_def,LIST_TYPE_def]) + >> (xapp>>xsimpl) +QED val r = translate spt_to_vecTheory.prepend_def; val r = translate spt_to_vecTheory.to_flat_def; @@ -1862,16 +1919,24 @@ val r = translate spt_to_vecTheory.spt_to_vec_def; val res = translate fromAList_def; val res = translate npbc_checkTheory.mk_subst_def; +(* +Definition print_lno_mini_def: + print_lno_mini (lno:num) mini = + case mini of NONE => toString lno ^ strlit " INF\n" + | SOME (i:num) => toString lno ^ strlit" " ^ toString i ^ strlit "\n" +End + +val res = translate print_lno_mini_def; *) + val check_red_arr = process_topdecs` - fun check_red_arr lno ord obj b tcb fml inds id c s pfs idopt earliest = + fun check_red_arr lno ord obj b tcb fml inds id c s pfs idopt vimap = let val s = mk_subst s in case red_fast s idopt pfs of None => ( - let val bortcb = b orelse tcb - val mini = get_earliest earliest s in - case reindex_partial_arr bortcb fml mini inds of - (rinds,(fmlls,rest)) => + let val bortcb = b orelse tcb in + case get_indices_arr bortcb fml inds s vimap of + (rinds,fmlls) => let val nc = not_1 c val rsubs = do_rso ord s c obj @@ -1884,20 +1949,25 @@ val check_red_arr = process_topdecs` let val u = rollback_arr fml' id id' val goals = subst_indexes_arr s bortcb fml' rinds in case red_cond_check fmlls nc pfs rsubs goals - of None => (fml', (rinds @ rest,id')) + of None => + (case set_indices inds s vimap rinds of + (inds',vimap') => + (fml', (inds', (vimap',id')))) | Some err => raise Fail (format_failure_2 lno ("Redundancy subproofs did not cover all subgoals. Info: " ^ err ^ ".") (print_subproofs_err rsubs goals)) end | Some cid => if check_contradiction_fml_arr b fml' cid then let val u = rollback_arr fml' id id' in - (fml', (rinds @ rest, id')) + (case set_indices inds s vimap rinds of + (inds',vimap') => + (fml', (inds', (vimap',id')))) end else raise Fail (format_failure lno ("did not derive contradiction from index:" ^ Int.toString cid))) end end) | Some (pf,cid) => - check_red_arr_fast lno b fml inds id c pf cid + check_red_arr_fast lno b fml inds id c pf cid vimap end` |> append_prog; (* Overloads all the _TYPEs that we will reuse *) @@ -1911,6 +1981,9 @@ Overload "ord_TYPE" = `` Overload "obj_TYPE" = `` OPTION_TYPE (PAIR_TYPE (LIST_TYPE (PAIR_TYPE INT NUM)) INT)`` +Overload "vimap_TYPE" = `` + SPTREE_SPT_TYPE (LIST_TYPE NUM)`` + Theorem check_red_arr_fast_spec: NUM lno lnov ∧ BOOL b bv ∧ @@ -1919,12 +1992,13 @@ Theorem check_red_arr_fast_spec: NUM id idv ∧ constraint_TYPE c cv ∧ LIST_TYPE NPBC_CHECK_LSTEP_TYPE pfs pfsv ∧ - NUM cid cidv + NUM cid cidv ∧ + vimap_TYPE vimap vimapv ⇒ app (p : 'ffi ffi_proj) ^(fetch_v "check_red_arr_fast" (get_ml_prog_state())) [lnov; bv; fmlv; indsv; idv; - cv; pfsv; cidv] + cv; pfsv; cidv; vimapv] (ARRAY fmlv fmllsv) (POSTve (λv. @@ -1932,18 +2006,20 @@ Theorem check_red_arr_fast_spec: ARRAY fmlv' fmllsv' * &( case check_red_list_fast b fmlls inds id - c pfs cid of NONE => F + c pfs cid vimap of NONE => F | SOME res => PAIR_TYPE (λl v. LIST_REL (OPTION_TYPE bconstraint_TYPE) l fmllsv' ∧ - v = fmlv') (PAIR_TYPE (LIST_TYPE NUM) NUM) res v + v = fmlv') + (PAIR_TYPE (LIST_TYPE NUM) + (PAIR_TYPE vimap_TYPE NUM)) res v )) (λe. SEP_EXISTS fmlv' fmllsv'. ARRAY fmlv' fmllsv' * & (Fail_exn e ∧ check_red_list_fast b fmlls inds id - c pfs cid = NONE))) + c pfs cid vimap = NONE))) Proof rw[]>> xcf "check_red_arr_fast" (get_ml_prog_state ())>> @@ -1992,12 +2068,12 @@ Theorem check_red_arr_spec: subst_raw_TYPE s sv ∧ pfs_TYPE pfs pfsv ∧ OPTION_TYPE NUM idopt idoptv ∧ - SPTREE_SPT_TYPE NUM earliest earliestv + vimap_TYPE vimap vimapv ⇒ app (p : 'ffi ffi_proj) ^(fetch_v "check_red_arr" (get_ml_prog_state())) [lnov; ordv; objv; bv; tcbv; fmlv; indsv; idv; - cv; sv; pfsv; idoptv; earliestv] + cv; sv; pfsv; idoptv; vimapv] (ARRAY fmlv fmllsv) (POSTve (λv. @@ -2005,18 +2081,20 @@ Theorem check_red_arr_spec: ARRAY fmlv' fmllsv' * &( case check_red_list ord obj b tcb fmlls inds id - c s pfs idopt earliest of NONE => F + c s pfs idopt vimap of NONE => F | SOME res => PAIR_TYPE (λl v. LIST_REL (OPTION_TYPE bconstraint_TYPE) l fmllsv' ∧ - v = fmlv') (PAIR_TYPE (LIST_TYPE NUM) NUM) res v + v = fmlv') + (PAIR_TYPE (LIST_TYPE NUM) + (PAIR_TYPE vimap_TYPE NUM)) res v )) (λe. SEP_EXISTS fmlv' fmllsv'. ARRAY fmlv' fmllsv' * & (Fail_exn e ∧ check_red_list ord obj b tcb fmlls inds id - c s pfs idopt earliest = NONE))) + c s pfs idopt vimap = NONE))) Proof rw[]>> xcf "check_red_arr" (get_ml_prog_state ())>> @@ -2040,7 +2118,6 @@ Proof xvar>>xsimpl)>> pairarg_tac>>gs[]>> xlet_autop>> - xlet_autop>> fs[PAIR_TYPE_def]>> xmatch>> rpt xlet_autop>> @@ -2099,7 +2176,9 @@ Proof fs[red_cond_check_def]>> pairarg_tac>>fs[]>> xlet_autop>> - xlet_autop>> + fs[PAIR_TYPE_def]>> + xmatch>> + rpt xlet_autop>> xcon>>xsimpl>> simp[PAIR_TYPE_def]>> xsimpl) @@ -2120,35 +2199,39 @@ Proof rw[]>>fs[]>> metis_tac[ARRAY_refl,NOT_EVERY,Fail_exn_def]) >> rpt xlet_autop>> + pairarg_tac>> + fs[PAIR_TYPE_def]>> + xmatch>> + rpt xlet_autop>> xcon>>xsimpl>> simp[PAIR_TYPE_def]>> xsimpl QED -val res = translate min_opt_def; -val res = translate update_earliest_def; +val res = translate opt_cons_def; +val res = translate update_vimap_def; val check_sstep_arr = process_topdecs` - fun check_sstep_arr lno step ord obj tcb fml inds id earliest = + fun check_sstep_arr lno step ord obj tcb fml inds id vimap = case step of Lstep lstep => (case check_lstep_arr lno lstep False fml 0 id of (rfml,(c,id')) => (case c of - None => (rfml,(inds,(earliest,id'))) + None => (rfml,(inds,(vimap,id'))) | Some cc => (Array.updateResize rfml None id' (Some cc), (sorted_insert id' inds, - (update_earliest earliest id' (fst (fst cc)), + (update_vimap vimap id' (fst (fst cc)), id'+1))) )) | Red c s pfs idopt => (case check_red_arr lno ord obj False tcb - fml inds id c s pfs idopt earliest of - (fml',(rinds,id')) => + fml inds id c s pfs idopt vimap of + (fml',(rinds,(vimap',id'))) => (Array.updateResize fml' None id' (Some (c,tcb)), (sorted_insert id' rinds, - (update_earliest earliest id' (fst c), + (update_vimap vimap' id' (fst c), id'+1)))) ` |> append_prog @@ -2163,11 +2246,11 @@ Theorem check_sstep_arr_spec: LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ (LIST_TYPE NUM) inds indsv ∧ NUM id idv ∧ - SPTREE_SPT_TYPE NUM earliest earliestv + vimap_TYPE vimap vimapv ⇒ app (p : 'ffi ffi_proj) ^(fetch_v "check_sstep_arr" (get_ml_prog_state())) - [lnov; stepv; ordv; objv; tcbv; fmlv; indsv; idv; earliestv] + [lnov; stepv; ordv; objv; tcbv; fmlv; indsv; idv; vimapv] (ARRAY fmlv fmllsv) (POSTve (λv. @@ -2175,20 +2258,20 @@ Theorem check_sstep_arr_spec: ARRAY fmlv' fmllsv' * &( case check_sstep_list step ord obj tcb - fmlls inds id earliest of NONE => F + fmlls inds id vimap of NONE => F | SOME res => PAIR_TYPE (λl v. LIST_REL (OPTION_TYPE bconstraint_TYPE) l fmllsv' ∧ v = fmlv') (PAIR_TYPE (LIST_TYPE NUM) - (PAIR_TYPE (SPTREE_SPT_TYPE NUM) NUM)) res v) + (PAIR_TYPE (vimap_TYPE) NUM)) res v) ) (λe. SEP_EXISTS fmlv' fmllsv'. ARRAY fmlv' fmllsv' * & (Fail_exn e ∧ check_sstep_list step ord obj tcb - fmlls inds id earliest = NONE))) + fmlls inds id vimap = NONE))) Proof rw[]>> xcf "check_sstep_arr" (get_ml_prog_state ())>> @@ -2247,25 +2330,26 @@ Proof xmatch>> xlet_autop>> rename1`check_red_list ord obj F tcb fmlls inds id - c s pfs idopt earliest`>> + c s pfs idopt vimap`>> xlet`POSTve (λv. SEP_EXISTS fmlv' fmllsv'. ARRAY fmlv' fmllsv' * &( case check_red_list ord obj F tcb fmlls inds id - c s pfs idopt earliest of NONE => F + c s pfs idopt vimap of NONE => F | SOME res => PAIR_TYPE (λl v. LIST_REL (OPTION_TYPE bconstraint_TYPE) l fmllsv' ∧ - v = fmlv') (PAIR_TYPE (LIST_TYPE NUM) NUM) res v + v = fmlv') (PAIR_TYPE (LIST_TYPE NUM) + (PAIR_TYPE vimap_TYPE NUM)) res v )) (λe. SEP_EXISTS fmlv' fmllsv'. ARRAY fmlv' fmllsv' * & (Fail_exn e ∧ check_red_list ord obj F tcb fmlls inds id - c s pfs idopt earliest = NONE))` + c s pfs idopt vimap = NONE))` >- ( xapp >> xsimpl>> CONJ_TAC >- EVAL_TAC>> @@ -3039,7 +3123,7 @@ val res = translate npbc_checkTheory.eq_obj_def; val res = translate npbc_checkTheory.check_eq_obj_def; val check_cstep_arr = process_topdecs` - fun check_cstep_arr lno cstep fml inds earliest pc = + fun check_cstep_arr lno cstep fml inds vimap pc = case cstep of Dom c s pfs idopt => ( case get_ord pc of None => @@ -3050,12 +3134,12 @@ val check_cstep_arr = process_topdecs` (fml',(rinds,id')) => (Array.updateResize fml' None id' (Some (c,get_tcb pc)), (sorted_insert id' rinds, - (update_earliest earliest id' (fst c), set_id pc (id'+1))))) + (update_vimap vimap id' (fst c), set_id pc (id'+1))))) | Sstep sstep => ( case check_sstep_arr lno sstep (get_ord pc) (get_obj pc) - (get_tcb pc) fml inds (get_id pc) earliest of - (fml',(inds',(earliest',id'))) => - (fml',(inds',(earliest',set_id pc id')))) + (get_tcb pc) fml inds (get_id pc) vimap of + (fml',(inds',(vimap',id'))) => + (fml',(inds',(vimap',set_id pc id')))) | Checkeddelete n s pfs idopt => ( if check_tcb_idopt_pc pc idopt then @@ -3066,35 +3150,35 @@ val check_cstep_arr = process_topdecs` case check_red_arr lno (get_ord pc) (get_obj pc) True (get_tcb pc) fml inds (get_id pc) - c s pfs idopt earliest of (fml',(inds',id')) => - (fml',(inds',(earliest,set_id pc id'))))) + c s pfs idopt vimap of (fml',(inds',(vimap',id'))) => + (fml',(inds',(vimap',set_id pc id'))))) else raise Fail (format_failure lno "invalid proof state for checked deletion")) | Uncheckeddelete ls => ( if check_tcb_ord pc then (list_delete_arr ls fml; - (fml, (inds, (earliest, set_chk pc False)))) + (fml, (inds, (vimap, set_chk pc False)))) else case all_core_arr fml inds [] of None => raise Fail (format_failure lno "not all constraints in core") | Some inds' => (list_delete_arr ls fml; - (fml, (inds', (earliest, set_chk pc False))))) + (fml, (inds', (vimap, set_chk pc False))))) | Transfer ls => ( let val fml' = core_from_inds_arr lno fml ls in - (fml', (inds, (earliest, pc))) + (fml', (inds, (vimap, pc))) end) | Strengthentocore b => ( let val inds' = fst (reindex_arr False fml inds) in if b then ( let val fml' = core_from_inds_arr lno fml inds' in - (fml', (inds', (earliest, set_tcb pc b))) + (fml', (inds', (vimap, set_tcb pc b))) end) else - (fml,(inds', (earliest, set_tcb pc b))) + (fml,(inds', (vimap, set_tcb pc b))) end) | Loadorder nn xs => let val inds' = fst (reindex_arr False fml inds) in @@ -3104,7 +3188,7 @@ val check_cstep_arr = process_topdecs` | Some ord' => if List.length xs = List.length (fst (snd ord')) then let val fml' = core_from_inds_arr lno fml inds' in - (fml', (inds', (earliest, set_ord pc (Some (ord',xs))))) + (fml', (inds', (vimap, set_ord pc (Some (ord',xs))))) end else raise Fail @@ -3114,11 +3198,11 @@ val check_cstep_arr = process_topdecs` (case get_ord pc of None => raise Fail (format_failure lno ("no order loaded")) | Some spo => - (fml,(inds,(earliest,set_ord pc None)))) + (fml,(inds,(vimap,set_ord pc None)))) | Storeorder name spo ws pfsr pfst => if check_storeorder spo ws pfst pfsr then - (fml, (inds, (earliest, set_orders pc ((name,spo)::get_orders pc)))) + (fml, (inds, (vimap, set_orders pc ((name,spo)::get_orders pc)))) else raise Fail (format_failure lno ("invalid order definition")) | Obj w mi bopt => ( @@ -3138,21 +3222,21 @@ val check_cstep_arr = process_topdecs` val c = model_improving obj new in (Array.updateResize fml None id (Some (c,True)), (sorted_insert id inds, - (update_earliest earliest id (fst c), + (update_vimap vimap id (fst c), obj_update pc (id+1) bound' dbound'))) end else - (fml, (inds, (earliest, obj_update pc (get_id pc) bound' dbound'))) + (fml, (inds, (vimap, obj_update pc (get_id pc) bound' dbound'))) end ) | Changeobj b fc' pfs => (case check_change_obj_arr lno b fml (get_id pc) (get_obj pc) fc' pfs of (fml',(fc',id')) => - (fml', (inds, (earliest, change_obj_update pc id' fc')))) + (fml', (inds, (vimap, change_obj_update pc id' fc')))) | Checkobj fc' => if check_eq_obj (get_obj pc) fc' - then (fml, (inds, (earliest, pc))) + then (fml, (inds, (vimap, pc))) else raise Fail (format_failure lno (err_obj_check_string (get_obj pc) fc')) @@ -3186,25 +3270,25 @@ Theorem check_cstep_arr_spec: LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ (LIST_TYPE NUM) inds indsv ∧ NPBC_CHECK_PROOF_CONF_TYPE pc pcv ∧ - SPTREE_SPT_TYPE NUM earliest earliestv + vimap_TYPE vimap vimapv ⇒ app (p : 'ffi ffi_proj) ^(fetch_v "check_cstep_arr" (get_ml_prog_state())) - [lnov; cstepv; fmlv; indsv; earliestv; pcv] + [lnov; cstepv; fmlv; indsv; vimapv; pcv] (ARRAY fmlv fmllsv) (POSTve (λv. SEP_EXISTS fmlv' fmllsv'. ARRAY fmlv' fmllsv' * &( - case check_cstep_list cstep fmlls inds earliest pc of + case check_cstep_list cstep fmlls inds vimap pc of NONE => F | SOME res => PAIR_TYPE (λl v. LIST_REL (OPTION_TYPE bconstraint_TYPE) l fmllsv' ∧ v = fmlv') (PAIR_TYPE (LIST_TYPE NUM) - (PAIR_TYPE (SPTREE_SPT_TYPE NUM) + (PAIR_TYPE (vimap_TYPE) NPBC_CHECK_PROOF_CONF_TYPE)) res v )) @@ -3212,7 +3296,7 @@ Theorem check_cstep_arr_spec: SEP_EXISTS fmlv' fmllsv'. ARRAY fmlv' fmllsv' * & (Fail_exn e ∧ - check_cstep_list cstep fmlls inds earliest pc = NONE))) + check_cstep_list cstep fmlls inds vimap pc = NONE))) Proof rw[]>> xcf "check_cstep_arr" (get_ml_prog_state ())>> @@ -3298,25 +3382,26 @@ Proof rpt xlet_autop>> fs[get_id_def,get_tcb_def,get_obj_def,get_ord_def]>> rename1`check_red_list ord obj T pc.tcb (delete_list n fmlls) - inds pc.id c s pfs idopt earliest`>> + inds pc.id c s pfs idopt vimap`>> xlet`POSTve (λv. SEP_EXISTS fmlv' fmllsv'. ARRAY fmlv' fmllsv' * &( case check_red_list ord obj T pc.tcb (delete_list n fmlls) - inds pc.id c s pfs idopt earliest of NONE => F + inds pc.id c s pfs idopt vimap of NONE => F | SOME res => PAIR_TYPE (λl v. LIST_REL (OPTION_TYPE bconstraint_TYPE) l fmllsv' ∧ - v = fmlv') (PAIR_TYPE (LIST_TYPE NUM) NUM) res v + v = fmlv') (PAIR_TYPE (LIST_TYPE NUM) + (PAIR_TYPE vimap_TYPE NUM)) res v )) (λe. SEP_EXISTS fmlv' fmllsv'. ARRAY fmlv' fmllsv' * & (Fail_exn e ∧ check_red_list ord obj T pc.tcb (delete_list n fmlls) - inds pc.id c s pfs idopt earliest = NONE))` + inds pc.id c s pfs idopt vimap = NONE))` >- ( xapp>>xsimpl>> CONJ_TAC >- EVAL_TAC>> diff --git a/examples/pseudo_bool/array/npbc_parseProgScript.sml b/examples/pseudo_bool/array/npbc_parseProgScript.sml index eb6bf6f4a6..730fc577aa 100644 --- a/examples/pseudo_bool/array/npbc_parseProgScript.sml +++ b/examples/pseudo_bool/array/npbc_parseProgScript.sml @@ -2312,15 +2312,15 @@ QED (* returns the necessary information to check the output and conclusion sections *) val check_unsat'' = process_topdecs ` - fun check_unsat'' fns fd lno fml inds earliest pc = + fun check_unsat'' fns fd lno fml inds vimap pc = case parse_cstep fns fd lno of (Inl s, (fns', lno')) => (lno', (s, (fns', (fml, (inds, pc))))) | (Inr cstep, (fns', lno')) => - (case check_cstep_arr lno cstep fml inds earliest pc of - (fml', (inds', (earliest', pc'))) => - check_unsat'' fns' fd lno' fml' inds' earliest' pc')` |> append_prog + (case check_cstep_arr lno cstep fml inds vimap pc of + (fml', (inds', (vimap', pc'))) => + check_unsat'' fns' fd lno' fml' inds' vimap' pc')` |> append_prog Theorem parse_sstep_LENGTH: ∀f ss res f' ss'. @@ -2365,15 +2365,15 @@ QED returning the last encountered state *) Definition parse_and_run_def: parse_and_run fns ss - fml inds earliest pc = + fml inds vimap pc = case parse_cstep fns ss of NONE => NONE | SOME (INL s, fns', rest) => SOME (rest, s, fns', fml, inds, pc) | SOME (INR cstep, fns', rest) => - (case check_cstep_list cstep fml inds earliest pc of - SOME (fml', inds', earliest', pc') => - parse_and_run fns' rest fml' inds' earliest' pc' + (case check_cstep_list cstep fml inds vimap pc of + SOME (fml', inds', vimap', pc') => + parse_and_run fns' rest fml' inds' vimap' pc' | res => NONE) Termination WF_REL_TAC `measure (LENGTH o FST o SND)`>> @@ -2412,19 +2412,19 @@ Proof QED Theorem check_unsat''_spec: - ∀fns ss fmlls inds earliest pc - fnsv lno lnov fmllsv indsv pcv lines fs fmlv earliestv. + ∀fns ss fmlls inds vimap pc + fnsv lno lnov fmllsv indsv pcv lines fs fmlv vimapv. fns_TYPE a fns fnsv ∧ NUM lno lnov ∧ LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ (LIST_TYPE NUM) inds indsv ∧ NPBC_CHECK_PROOF_CONF_TYPE pc pcv ∧ - SPTREE_SPT_TYPE NUM earliest earliestv ∧ + vimap_TYPE vimap vimapv ∧ MAP toks_fast lines = ss ⇒ app (p : 'ffi ffi_proj) ^(fetch_v "check_unsat''" (get_ml_prog_state())) - [fnsv; fdv; lnov; fmlv; indsv; earliestv; pcv] + [fnsv; fdv; lnov; fmlv; indsv; vimapv; pcv] (STDIO fs * INSTREAM_LINES fd fdv lines fs * ARRAY fmlv fmllsv) (POSTve (λv. @@ -2433,7 +2433,7 @@ Theorem check_unsat''_spec: INSTREAM_LINES fd fdv lines' (forwardFD fs fd k) * ARRAY fmlv' fmllsv' * &( - parse_and_run fns ss fmlls inds earliest pc = + parse_and_run fns ss fmlls inds vimap pc = SOME (MAP toks_fast lines',res) ∧ PAIR_TYPE NUM ( PAIR_TYPE (LIST_TYPE (SUM_TYPE STRING_TYPE INT)) ( @@ -2448,7 +2448,7 @@ Theorem check_unsat''_spec: ARRAY fmlv' fmllsv' * STDIO (forwardFD fs fd k) * INSTREAM_LINES fd fdv lines' (forwardFD fs fd k) * &(Fail_exn e ∧ - parse_and_run fns ss fmlls inds earliest pc = NONE))) + parse_and_run fns ss fmlls inds vimap pc = NONE))) Proof ho_match_mp_tac (fetch "-" "parse_and_run_ind")>> rw[]>> @@ -2626,37 +2626,35 @@ QED val _ = translate rev_enum_full_def; -val res = translate npbc_listTheory.update_earliest_def; - -Definition fold_update_earliest_enum_def: - (fold_update_earliest_enum k [] acc = acc) ∧ - (fold_update_earliest_enum k (x::xs) acc = - fold_update_earliest_enum (k+1) - xs (update_earliest acc k (FST x))) +Definition fold_update_vimap_enum_def: + (fold_update_vimap_enum (k:num) [] acc = acc) ∧ + (fold_update_vimap_enum k (x::xs) acc = + fold_update_vimap_enum (k+1) + xs (update_vimap acc k (FST x))) End -Definition fold_update_earliest_enum_full_def: - fold_update_earliest_enum_full k fml = - fold_update_earliest_enum k fml LN +Definition fold_update_vimap_enum_full_def: + fold_update_vimap_enum_full k fml = + fold_update_vimap_enum k fml LN End -Theorem fold_update_earliest_enum_FOLDL: +Theorem fold_update_vimap_enum_FOLDL: ∀xs k acc. - fold_update_earliest_enum k xs acc = - (FOLDL (λacc (i,v). update_earliest acc i (FST v)) acc (enumerate k xs)) + fold_update_vimap_enum k xs acc = + (FOLDL (λacc (i,v). update_vimap acc i (FST v)) acc (enumerate k xs)) Proof - Induct>>rw[fold_update_earliest_enum_def,miscTheory.enumerate_def] + Induct>>rw[fold_update_vimap_enum_def,miscTheory.enumerate_def] QED -Theorem fold_update_earliest_enum_full_FOLDL: - fold_update_earliest_enum_full k xs = - (FOLDL (λacc (i,v). update_earliest acc i (FST v)) LN (enumerate k xs)) +Theorem fold_update_vimap_enum_full_FOLDL: + fold_update_vimap_enum_full k xs = + (FOLDL (λacc (i,v). update_vimap acc i (FST v)) LN (enumerate k xs)) Proof - rw[fold_update_earliest_enum_full_def,fold_update_earliest_enum_FOLDL] + rw[fold_update_vimap_enum_full_def,fold_update_vimap_enum_FOLDL] QED -val res = translate fold_update_earliest_enum_def; -val res = translate fold_update_earliest_enum_full_def; +val res = translate fold_update_vimap_enum_def; +val res = translate fold_update_vimap_enum_full_def; val res = translate parse_unsat_def; @@ -2936,10 +2934,10 @@ val check_unsat' = process_topdecs ` val arr = Array.array (2*id) None val arr = fill_arr arr 1 fml val inds = rev_enum_full 1 fml - val earliest = fold_update_earliest_enum_full 1 fml + val vimap = fold_update_vimap_enum_full 1 fml val pc = init_conf id True obj in - (case check_unsat'' fns fd lno arr inds earliest pc of + (case check_unsat'' fns fd lno arr inds vimap pc of (lno', (s, (fns',( (fml', (inds', pc')))))) => conv_boutput_hconcl @@ -2952,10 +2950,10 @@ val check_unsat' = process_topdecs ` end` |> append_prog; Theorem parse_and_run_check_csteps_list: - ∀fns ss fml inds earliest pc rest s fns' fml' inds' pc'. - parse_and_run fns ss fml inds earliest pc = SOME (rest, s, fns', (fml', inds', pc')) ⇒ - ∃csteps earliest'. - check_csteps_list csteps fml inds earliest pc = SOME (fml', inds', earliest', pc') + ∀fns ss fml inds vimap pc rest s fns' fml' inds' pc'. + parse_and_run fns ss fml inds vimap pc = SOME (rest, s, fns', (fml', inds', pc')) ⇒ + ∃csteps vimap'. + check_csteps_list csteps fml inds vimap pc = SOME (fml', inds', vimap', pc') Proof ho_match_mp_tac parse_and_run_ind>> rw[]>> @@ -3048,9 +3046,9 @@ Proof `BOOL T (Conv (SOME (TypeStamp "True" 0)) [])` by EVAL_TAC>> xlet_autop>> - qmatch_asmsub_abbrev_tac`SPTREE_SPT_TYPE NUM earliest earliestv`>> + qmatch_asmsub_abbrev_tac`vimap_TYPE vimap vimapv`>> Cases_on` - parse_and_run fns (MAP toks_fast lines) fmlls inds earliest + parse_and_run fns (MAP toks_fast lines) fmlls inds vimap (init_conf (LENGTH fml + 1) T obj)` >- ( (* fail to parse and run *) @@ -3105,7 +3103,7 @@ Proof ARRAY fmlv' fmllsv' * &( parse_and_run fns (MAP toks_fast lines) - fmlls inds earliest (init_conf (LENGTH fml + 1) T obj) = + fmlls inds vimap (init_conf (LENGTH fml + 1) T obj) = SOME (MAP toks_fast lines',res) ∧ PAIR_TYPE NUM ( PAIR_TYPE (LIST_TYPE (SUM_TYPE STRING_TYPE INT)) ( @@ -3191,13 +3189,13 @@ Proof match_mp_tac (GEN_ALL npbc_listTheory.check_csteps_list_concl)>> first_x_assum (irule_at Any)>> unabbrev_all_tac>> - gs[rev_enum_full_rev_enumerate, fold_update_earliest_enum_full_FOLDL]>> + gs[rev_enum_full_rev_enumerate, fold_update_vimap_enum_full_FOLDL]>> metis_tac[])>> simp[get_bound_def]>> match_mp_tac (GEN_ALL npbc_listTheory.check_csteps_list_output)>> first_x_assum (irule_at Any)>> unabbrev_all_tac>> - gs[rev_enum_full_rev_enumerate, fold_update_earliest_enum_full_FOLDL]>> + gs[rev_enum_full_rev_enumerate, fold_update_vimap_enum_full_FOLDL]>> metis_tac[])>> metis_tac[STDIO_INSTREAM_LINES_refl_gc])>> xsimpl From 8a7bef803a13dfc36e826febf6d3b1a67c9cd898 Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Fri, 26 Jan 2024 20:40:44 +0800 Subject: [PATCH 12/38] refactor reindexing --- .../pseudo_bool/array/npbc_listScript.sml | 188 +++++++++++------- 1 file changed, 117 insertions(+), 71 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 426ffe56e4..2fb6b89ff3 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -537,6 +537,7 @@ Definition check_subproofs_list_def: | _ => NONE) End +(* Definition reindex_aux_def: (reindex_aux b fml [] iacc vacc = (REVERSE iacc, vacc)) ∧ @@ -553,7 +554,37 @@ End Definition reindex_def: (reindex b fml is = reindex_aux b fml is [] []) End +*) + +Definition reindex_aux_def: + (reindex_aux fml [] iacc = REVERSE iacc) ∧ + (reindex_aux fml (i::is) iacc = + case any_el i fml NONE of + NONE => reindex_aux fml is iacc + | SOME (v,b') => + reindex_aux fml is (i::iacc)) +End + +Definition reindex_def: + (reindex fml is = reindex_aux fml is []) +End + +Definition revalue_aux_def: + (revalue_aux b fml [] vacc = vacc) ∧ + (revalue_aux b fml (i::is) vacc = + case any_el i fml NONE of + NONE => revalue_aux b fml is vacc + | SOME (v,b') => + let vacc' = + if b ⇒ b' then v::vacc else vacc in + revalue_aux b fml is vacc') +End +Definition revalue_def: + (revalue b fml is = revalue_aux b fml is []) +End + +(* Definition reindex_partial_aux_def: (reindex_partial_aux b fml mini [] iacc vacc = (REVERSE iacc, vacc,[])) ∧ @@ -574,6 +605,7 @@ Definition reindex_partial_def: | SOME mini => reindex_partial_aux b fml mini is [] [] End +*) Definition subst_opt_subst_fun_def: subst_opt_subst_fun s c = subst_opt (subst_fun s) c @@ -651,30 +683,6 @@ Proof every_case_tac>>fs[] QED -Definition split_goals_hash_def: - split_goals_hash fmlls extra (proved:num_set) - (goals:(num # (int # num) list # num) list) = - let (lp,lf) = - PARTITION (λ(i,c). lookup i proved ≠ NONE) goals in - let lf = FILTER (λc. ¬(imp extra c)) (MAP SND lf) in - let proved = MAP SND lp in - let hs = mk_hashset fmlls (mk_hashset proved (REPLICATE splim [])) in - EVERY (λc. in_hashset c hs) lf -End - -(* Not meant to be executed, mainly just abbrevation... *) -Definition do_red_check_def: - do_red_check idopt b tcb fml - s rfml rinds fmlls extra pfs rsubs = - case idopt of NONE => - let goals = subst_indexes s (b ∨ tcb) rfml rinds in - let (l,r) = extract_pids pfs LN LN in - split_goals_hash fmlls extra l goals ∧ - EVERY (λid. lookup id r ≠ NONE) (COUNT_LIST (LENGTH rsubs)) - | SOME cid => - check_contradiction_fml_list b fml cid -End - (* Fast path for a special case *) Definition red_fast_def: red_fast s idopt pfs = ( @@ -703,6 +711,33 @@ Definition check_red_list_fast_def: NONE End +(* The fmlls argument should be delayed and avoided + as much as possible *) +Definition split_goals_hash_def: + split_goals_hash fmlls extra (proved:num_set) + (goals:(num # (int # num) list # num) list) = + let (lp,lf) = + PARTITION (λ(i,c). lookup i proved ≠ NONE) goals in + let lf = FILTER (λc. ¬(imp extra c)) (MAP SND lf) in + let proved = MAP SND lp in + let hs = mk_hashset fmlls (mk_hashset proved (REPLICATE splim [])) in + EVERY (λc. in_hashset c hs) lf +End + +(* Not meant to be executed, mainly just abbrevation... *) +Definition do_red_check_def: + do_red_check idopt b tcb fml inds + s rfml rinds extra pfs rsubs = + case idopt of NONE => + let goals = subst_indexes s (b ∨ tcb) rfml rinds in + let (l,r) = extract_pids pfs LN LN in + let fmlls = revalue (b ∨ tcb) rfml inds in + split_goals_hash fmlls extra l goals ∧ + EVERY (λid. lookup id r ≠ NONE) (COUNT_LIST (LENGTH rsubs)) + | SOME cid => + check_contradiction_fml_list b fml cid +End + (* Definition get_earliest_def: (get_earliest earliest (INR v) = @@ -714,15 +749,15 @@ End *) (* A reverse mapping of vars -> indices *) Definition get_indices_def: - get_indices b tcb fml inds s vimap = + get_indices fml inds s vimap = case s of INR v => - if length v = 0 then ([],[]) - else (reindex (b ∨ tcb) fml inds) + if length v = 0 then [] + else reindex fml inds | INL (n,_) => case sptree$lookup n vimap of - NONE => ([],[]) - | SOME inds => (reindex (b ∨ tcb) fml inds) + NONE => [] + | SOME inds => reindex fml inds End Definition set_indices_def: @@ -740,7 +775,8 @@ Definition check_red_list_def: let s = mk_subst s in case red_fast s idopt pfs of NONE => ( - let (rinds,fmlls) = get_indices b tcb fml inds s vimap in + let rinds = get_indices fml inds s vimap in + let (inds',vimap') = set_indices inds s vimap rinds in let nc = not c in let fml_not_c = update_resize fml NONE (SOME (nc,b)) id in let rsubs = red_subgoals ord (subst_fun s) c obj in @@ -752,9 +788,8 @@ Definition check_red_list_def: NONE => NONE | SOME(fml', id') => let rfml = rollback fml' id id' in - if do_red_check idopt b tcb fml' - s rfml rinds fmlls nc pfs rsubs then - let (inds',vimap') = set_indices inds s vimap rinds in + if do_red_check idopt b tcb fml' inds' + s rfml rinds nc pfs rsubs then SOME (rfml,inds',vimap',id') else NONE)) | SOME (pf,cid) => @@ -1101,27 +1136,32 @@ Proof QED Theorem reindex_aux: - ∀inds iacc vacc. - reindex_aux b fmlls inds iacc vacc = + ∀inds iacc. + reindex_aux fmlls inds iacc = let is = FILTER (λx. IS_SOME (any_el x fmlls NONE)) inds in - let vs = - MAP (λx. THE (lookup_core_only_list b fmlls x)) - (FILTER (λx. IS_SOME (lookup_core_only_list b fmlls x)) - inds) in - (REVERSE iacc ++ is, REVERSE vs ++ vacc) + (REVERSE iacc ++ is) Proof Induct>>rw[reindex_aux_def]>> gvs[lookup_core_only_list_def,IS_SOME_EXISTS,AllCaseEqs()] QED -Theorem FST_reindex_characterize: - reindex b fmlls inds = (is,vs) ⇒ - is = FILTER (λx. IS_SOME (any_el x fmlls NONE)) inds +Theorem reindex_characterize: + reindex fmlls inds = FILTER (λx. IS_SOME (any_el x fmlls NONE)) inds Proof rw[reindex_def,reindex_aux] QED -Theorem SND_reindex_characterize: +Theorem revalue_aux_SUBSET: + ∀inds vacc. + fml_rel fml fmlls ∧ + set vacc ⊆ core_only_fml b fml ⇒ + set (revalue_aux b fmlls inds vacc) ⊆ core_only_fml b fml +Proof + cheat +QED + +(* +SND_reindex_characterize: fml_rel fml fmlls ∧ reindex b fmlls inds = (is,vs) ⇒ set vs ⊆ core_only_fml b fml @@ -1138,24 +1178,30 @@ Proof rw[]>>fs[]>> metis_tac[] QED +*) + +Theorem revalue_SUBSET: + fml_rel fml fmlls ==> + set (revalue b fmlls inds) ⊆ core_only_fml b fml +Proof + rw[revalue_def,revalue_aux_SUBSET] +QED Theorem SORTED_reindex: SORTED $>= inds ∧ - reindex b fml inds = (is,vs) ⇒ + reindex fml inds = is ⇒ SORTED $>= is Proof - rw[]>>drule FST_reindex_characterize>> - rw[]>> + rw[reindex_characterize]>> match_mp_tac SORTED_FILTER>> fs[transitive_def] QED Theorem ind_rel_reindex: - ind_rel fml inds ∧ - reindex b fml inds = (is,vs) ⇒ - ind_rel fml is + ind_rel fml inds ⇒ + ind_rel fml (reindex fml inds) Proof - rw[]>>drule FST_reindex_characterize>> + rw[reindex_characterize]>> fs[ind_rel_def,MEM_FILTER] QED @@ -1171,6 +1217,7 @@ Proof fs[] QED +(* Theorem reindex_partial_aux: ∀inds iacc vacc. SORTED $>= inds ⇒ @@ -1257,6 +1304,7 @@ Proof \\ rpt $ irule_at Any sortingTheory.SORTED_FILTER \\ gvs [] \\ gvs [transitive_def] QED +*) Theorem MEM_subst_indexes: ∀inds i c. @@ -1406,7 +1454,7 @@ Proof QED Theorem ind_rel_get_indices_set_indices: - get_indices b tcb fmlls inds s vimap = (rinds,_) ∧ + get_indices fmlls inds s vimap = rinds ∧ set_indices inds s vimap rinds = (inds',vimap') ∧ ind_rel fmlls inds ⇒ ind_rel fmlls inds' @@ -1417,8 +1465,8 @@ Proof QED Theorem vimap_rel_get_indices_set_indices: - get_indices b tcb fmlls inds s vimap = (rinds,_) ∧ - set_indices inds s vimap rinds = (inds',vimap') ∧ + set_indices inds s vimap + (get_indices fmlls inds s vimap) = (inds',vimap') ∧ vimap_rel fmlls vimap ⇒ vimap_rel fmlls vimap' Proof @@ -1446,7 +1494,6 @@ Proof fs[check_red_list_def]>> gvs[AllCaseEqs()] >- ( - pairarg_tac>>fs[]>> pairarg_tac>>fs[]>> every_case_tac>>gvs[]>> simp[check_red_def]>> @@ -1476,9 +1523,10 @@ Proof disch_then (qspec_then`mk_core_fml (b ∨ tcb) fml` mp_tac)>> impl_tac >- ( simp[range_mk_core_fml]>> - gvs[get_indices_def] >> every_case_tac>> gvs[]>> - match_mp_tac (GEN_ALL SND_reindex_characterize)>> - metis_tac[])>> + gvs[get_indices_def,set_indices_def] >> + every_case_tac>> gvs[]>> + match_mp_tac revalue_SUBSET>> + match_mp_tac fml_rel_rollback>>rw[]>>fs[])>> match_mp_tac split_goals_same_goals>> simp[EXTENSION,FORALL_PROD]>> rw[]>>eq_tac>>rw[] @@ -1868,14 +1916,15 @@ Proof QED Definition do_dom_check_def: - do_dom_check idopt fml rfml w indcore indfml extra pfs dsubs = + do_dom_check idopt fml rfml w indcore rinds extra pfs dsubs = case idopt of NONE => let goals = MAP_OPT (subst_opt w) indcore in let (l,r) = extract_pids pfs LN LN in if EVERY (λid. lookup id r ≠ NONE) (COUNT_LIST (LENGTH dsubs)) then - split_goals_hash indfml extra l goals + let fmlls = revalue F rfml rinds in + split_goals_hash fmlls extra l goals else F | SOME cid => check_contradiction_fml_list F fml cid @@ -1974,7 +2023,7 @@ Definition check_cstep_list_def: | SOME spo => ( let nc = not c in let id = pc.id in - let (rinds,fmlls) = reindex F fml inds in + let rinds = reindex fml inds in let corels = core_fmlls fml rinds in let fml_not_c = update_resize fml NONE (SOME (nc,F)) id in let s = mk_subst s in @@ -1987,7 +2036,7 @@ Definition check_cstep_list_def: NONE => NONE | SOME (fml',id') => let rfml = rollback fml' id id' in - if do_dom_check idopt fml' rfml w corels fmlls nc pfs dsubs then + if do_dom_check idopt fml' rfml w corels rinds nc pfs dsubs then SOME( update_resize rfml NONE (SOME (c,pc.tcb)) id', sorted_insert id' rinds, @@ -2028,7 +2077,7 @@ Definition check_cstep_list_def: | SOME fml' => SOME (fml', inds, vimap, pc)) | StrengthenToCore b => - (let inds' = FST (reindex F fml inds) in + (let inds' = reindex fml inds in let pc' = pc with tcb := b in if b then @@ -2037,7 +2086,7 @@ Definition check_cstep_list_def: else SOME (fml,inds',vimap,pc')) | LoadOrder nn xs => - (let inds' = FST (reindex F fml inds) in + (let inds' = reindex fml inds in case ALOOKUP pc.orders nn of NONE => NONE | SOME ord' => if LENGTH xs = LENGTH (FST (SND ord')) then @@ -2243,9 +2292,7 @@ Proof (drule_at Any) fml_rel_check_subproofs_list>> disch_then(qspec_then`insert pc.id (not p,F) fml` mp_tac)>> impl_tac >- ( - simp[fml_rel_update_resize,any_el_update_resize]>> - match_mp_tac ind_rel_update_resize_sorted_insert>> - metis_tac[ind_rel_reindex])>> + simp[fml_rel_update_resize,any_el_update_resize])>> rw[]>>simp[]>> drule check_subproofs_list_id>> drule check_subproofs_list_id_upper>> @@ -2260,8 +2307,9 @@ Proof (drule_at Any) split_goals_hash_imp_split_goals>> disch_then(qspec_then `mk_core_fml F fml` mp_tac)>> impl_tac >- ( - drule_all (GEN_ALL SND_reindex_characterize)>> - simp[range_mk_core_fml])>> + simp[range_mk_core_fml]>> + match_mp_tac revalue_SUBSET>> + match_mp_tac fml_rel_rollback>>rw[]>>fs[])>> match_mp_tac split_goals_same_goals>> simp[EXTENSION,FORALL_PROD,MEM_toAList,lookup_map_opt,MEM_MAP_OPT,AllCaseEqs(),lookup_mk_core_fml]>> simp[MEM_core_fmlls]>> @@ -2339,7 +2387,6 @@ Proof metis_tac[IS_SOME_EXISTS,option_CLAUSES]) >- ( (* StrengthenToCore *) gvs[check_cstep_list_def,AllCaseEqs(),check_cstep_def]>> - Cases_on`reindex F fmlls inds`>> drule_all ind_rel_reindex >- ( drule any_el_core_from_inds>> @@ -2372,7 +2419,6 @@ Proof fs[]) >- ( (* LoadOrder *) gvs[check_cstep_list_def,AllCaseEqs(),check_cstep_def]>> - Cases_on`reindex F fmlls inds`>> drule_all ind_rel_reindex>> drule any_el_core_from_inds>> strip_tac>>fs[]>> From f505e0eb3cf11b693f82463e8f9d2e37f5b57b1e Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Fri, 26 Jan 2024 22:56:41 +0800 Subject: [PATCH 13/38] fix in npbc_arrayProg --- .../array/npbc_arrayProgScript.sml | 275 +++++++++++------- 1 file changed, 169 insertions(+), 106 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_arrayProgScript.sml b/examples/pseudo_bool/array/npbc_arrayProgScript.sml index 5253bbe59d..2495a82428 100644 --- a/examples/pseudo_bool/array/npbc_arrayProgScript.sml +++ b/examples/pseudo_bool/array/npbc_arrayProgScript.sml @@ -931,6 +931,86 @@ QED Theorem check_lstep_arr_spec = CONJUNCT1 check_lstep_arr_spec_aux Theorem check_lsteps_arr_spec = CONJUNCT2 check_lstep_arr_spec_aux +val reindex_aux_arr = process_topdecs ` + fun reindex_aux_arr fml ls iacc = + case ls of + [] => (List.rev iacc) + | (i::is) => + case Array.lookup fml None i of + None => reindex_aux_arr fml is iacc + | Some vb => + reindex_aux_arr fml is (i::iacc)` |> append_prog; + +val reindex_arr = process_topdecs ` + fun reindex_arr fml is = + reindex_aux_arr fml is []` + |> append_prog; + +Theorem reindex_aux_arr_spec: + ∀inds indsv fmlls fmlv iacc iaccv. + LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ + (LIST_TYPE NUM) inds indsv ∧ + (LIST_TYPE NUM) iacc iaccv + ⇒ + app (p : 'ffi ffi_proj) + ^(fetch_v "reindex_aux_arr" (get_ml_prog_state())) + [fmlv; indsv; iaccv] + (ARRAY fmlv fmllsv) + (POSTv v. + ARRAY fmlv fmllsv * + &(LIST_TYPE NUM + (reindex_aux fmlls inds iacc) v)) +Proof + Induct>> + xcf"reindex_aux_arr"(get_ml_prog_state ())>> + fs[LIST_TYPE_def] + >- ( + xmatch>> + xapp_spec (ListProgTheory.reverse_v_thm |> INST_TYPE [alpha |-> ``:num``])>> + xsimpl>> + simp[reindex_aux_def,LIST_TYPE_def,PAIR_TYPE_def]>> + metis_tac[])>> + xmatch>> + xlet_auto>- (xcon>>xsimpl)>> + xlet_auto>> + `OPTION_TYPE bconstraint_TYPE + (any_el h fmlls NONE) v'` by ( + rw[any_el_ALT]>> + fs[LIST_REL_EL_EQN,OPTION_TYPE_def])>> + rw[]>> + simp[reindex_aux_def]>> + TOP_CASE_TAC>>fs[OPTION_TYPE_def] + >- ( + xmatch>> + xapp>>simp[])>> + xmatch>> + xlet_autop>> + xapp>> + fs[LIST_TYPE_def] +QED + +Theorem reindex_arr_spec: + ∀inds indsv fmlls fmlv. + LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ + (LIST_TYPE NUM) inds indsv + ⇒ + app (p : 'ffi ffi_proj) + ^(fetch_v "reindex_arr" (get_ml_prog_state())) + [fmlv; indsv] + (ARRAY fmlv fmllsv) + (POSTv v. + ARRAY fmlv fmllsv * + &(LIST_TYPE NUM + (reindex fmlls inds) v)) +Proof + xcf"reindex_arr"(get_ml_prog_state ())>> + rpt xlet_autop>> + xapp>> + xsimpl>> + simp[reindex_def]>> + metis_tac[LIST_TYPE_def] +QED + Definition mk_vacc_def: mk_vacc b b' v vacc = if b ⇒ b' then v::vacc else vacc @@ -938,47 +1018,45 @@ End val r = translate mk_vacc_def; -val reindex_aux_arr = process_topdecs ` - fun reindex_aux_arr b fml ls iacc vacc = +val revalue_aux_arr = process_topdecs ` + fun revalue_aux_arr b fml ls vacc = case ls of - [] => (List.rev iacc, vacc) + [] => vacc | (i::is) => case Array.lookup fml None i of - None => reindex_aux_arr b fml is iacc vacc + None => revalue_aux_arr b fml is vacc | Some (v,b') => - reindex_aux_arr b fml is (i::iacc) + revalue_aux_arr b fml is (mk_vacc b b' v vacc)` |> append_prog; -val reindex_arr = process_topdecs ` - fun reindex_arr b fml is = - reindex_aux_arr b fml is [] []` +val revalue_arr = process_topdecs ` + fun revalue_arr b fml is = + revalue_aux_arr b fml is []` |> append_prog; -Theorem reindex_aux_arr_spec: - ∀inds indsv b bv fmlls fmlv iacc iaccv vacc vaccv. +Theorem revalue_aux_arr_spec: + ∀inds indsv fmlls fmlv vacc vaccv. BOOL b bv ∧ LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ (LIST_TYPE NUM) inds indsv ∧ - (LIST_TYPE NUM) iacc iaccv ∧ (LIST_TYPE constraint_TYPE) vacc vaccv ⇒ app (p : 'ffi ffi_proj) - ^(fetch_v "reindex_aux_arr" (get_ml_prog_state())) - [bv; fmlv; indsv; iaccv; vaccv] + ^(fetch_v "revalue_aux_arr" (get_ml_prog_state())) + [bv; fmlv; indsv; vaccv] (ARRAY fmlv fmllsv) (POSTv v. ARRAY fmlv fmllsv * - &(PAIR_TYPE (LIST_TYPE NUM) (LIST_TYPE constraint_TYPE) - (reindex_aux b fmlls inds iacc vacc) v)) + &(LIST_TYPE constraint_TYPE + (revalue_aux b fmlls inds vacc) v)) Proof Induct>> - xcf"reindex_aux_arr"(get_ml_prog_state ())>> + xcf"revalue_aux_arr"(get_ml_prog_state ())>> fs[LIST_TYPE_def] >- ( xmatch>> - rpt xlet_autop>> - xcon>>xsimpl>> - simp[reindex_aux_def,LIST_TYPE_def,PAIR_TYPE_def])>> + xvar>>xsimpl>> + simp[revalue_aux_def,LIST_TYPE_def,PAIR_TYPE_def])>> xmatch>> xlet_auto>- (xcon>>xsimpl)>> xlet_auto>> @@ -987,7 +1065,7 @@ Proof rw[any_el_ALT]>> fs[LIST_REL_EL_EQN,OPTION_TYPE_def])>> rw[]>> - simp[reindex_aux_def]>> + simp[revalue_aux_def]>> TOP_CASE_TAC>>fs[OPTION_TYPE_def] >- ( xmatch>> @@ -999,30 +1077,30 @@ Proof fs[LIST_TYPE_def,mk_vacc_def] QED -Theorem reindex_arr_spec: +Theorem revalue_arr_spec: ∀inds indsv fmlls fmlv. BOOL b bv ∧ LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ (LIST_TYPE NUM) inds indsv ⇒ app (p : 'ffi ffi_proj) - ^(fetch_v "reindex_arr" (get_ml_prog_state())) + ^(fetch_v "revalue_arr" (get_ml_prog_state())) [bv; fmlv; indsv] (ARRAY fmlv fmllsv) (POSTv v. ARRAY fmlv fmllsv * - &(PAIR_TYPE (LIST_TYPE NUM) - (LIST_TYPE constraint_TYPE) - (reindex b fmlls inds) v)) + &(LIST_TYPE constraint_TYPE + (revalue b fmlls inds) v)) Proof - xcf"reindex_arr"(get_ml_prog_state ())>> + xcf"revalue_arr"(get_ml_prog_state ())>> rpt xlet_autop>> xapp>> xsimpl>> - simp[reindex_def]>> + simp[revalue_def]>> metis_tac[LIST_TYPE_def] QED +(* val reindex_partial_aux_arr = process_topdecs ` fun reindex_partial_aux_arr b fml mini ls iacc vacc = case ls of @@ -1129,6 +1207,7 @@ Proof simp[reindex_partial_def]>> metis_tac[LIST_TYPE_def] QED +*) val res = translate is_Pos_def; val res = translate subst_aux_def; @@ -1698,11 +1777,12 @@ Proof QED Definition red_cond_check_def: - red_cond_check indfml extra + red_cond_check bortcb fml inds extra (pfs:(( ((num + num) # num) option, (lstep list)) alist)) (rsubs:((int # num) list # num) list list) goals = let (l,r) = extract_pids pfs LN LN in - split_goals_hash indfml extra l goals ∧ + let fmlls = revalue bortcb fml inds in + split_goals_hash fmlls extra l goals ∧ EVERY (λid. lookup id r ≠ NONE) (COUNT_LIST (LENGTH rsubs)) End @@ -1723,11 +1803,12 @@ Definition red_cond_check_pure_def: End Theorem red_cond_check_eq: - red_cond_check indfml extra pfs rsubs goals = + red_cond_check bortcb fml inds extra pfs rsubs goals = case red_cond_check_pure extra pfs rsubs goals of NONE => F | SOME (x,ls) => - let hs = mk_hashset indfml (mk_hashset x (REPLICATE splim [])) in + let fmlls = revalue bortcb fml inds in + let hs = mk_hashset fmlls (mk_hashset x (REPLICATE splim [])) in EVERY (λc. in_hashset c hs) ls Proof rw[red_cond_check_def,red_cond_check_pure_def]>> @@ -1744,14 +1825,18 @@ val res = translate PARTITION_DEF; val res = translate red_cond_check_pure_def; val red_cond_check = process_topdecs` - fun red_cond_check indfml extra pfs rsubs goals = + fun red_cond_check bortcb fml inds extra pfs rsubs goals = case red_cond_check_pure extra pfs rsubs goals of None => Some "not all # subgoals present" | Some (x,ls) => - hash_check indfml x ls` |> append_prog + let val fmlls = revalue_arr bortcb fml inds in + hash_check fmlls x ls + end` |> append_prog Theorem red_cond_check_spec: - LIST_TYPE constraint_TYPE a av ∧ + BOOL bt btv ∧ + LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ + LIST_TYPE NUM inds indsv ∧ constraint_TYPE aa aav ∧ LIST_TYPE (PAIR_TYPE (OPTION_TYPE (PAIR_TYPE (SUM_TYPE NUM NUM) NUM)) @@ -1761,12 +1846,13 @@ Theorem red_cond_check_spec: ⇒ app (p : 'ffi ffi_proj) ^(fetch_v "red_cond_check" (get_ml_prog_state())) - [av; aav; bv; cv; dv] - emp + [btv; fmlv; indsv; aav; bv; cv; dv] + (ARRAY fmlv fmllsv) (POSTv resv. + ARRAY fmlv fmllsv * & ∃err. OPTION_TYPE STRING_TYPE - (if red_cond_check a aa b c d then NONE + (if red_cond_check bt fmlls inds aa b c d then NONE else SOME err) resv) Proof rw[]>> @@ -1781,6 +1867,7 @@ Proof xcon>>xsimpl)>> Cases_on`x`>>fs[PAIR_TYPE_def]>> xmatch>> + xlet_autop>> xapp>> xsimpl>> fs[]>> @@ -1870,14 +1957,15 @@ End val res = translate get_indices_pure_def; val get_indices_arr = process_topdecs` - fun get_indices_arr bortcb fml inds s vimap = + fun get_indices_arr fml inds s vimap = case get_indices_pure inds s vimap of - Inl v => ([],[]) - | Inr inds => - (reindex_arr bortcb fml inds)` |> append_prog; + Inl v => [] + | Inr inds => reindex_arr fml inds` |> append_prog; + +Overload "vimap_TYPE" = `` + SPTREE_SPT_TYPE (LIST_TYPE NUM)`` Theorem get_indices_arr_spec: - BOOL (b ∨ tcb) bv ∧ LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ (LIST_TYPE NUM) inds indsv ∧ subst_TYPE s sv ∧ @@ -1885,14 +1973,13 @@ Theorem get_indices_arr_spec: ⇒ app (p : 'ffi ffi_proj) ^(fetch_v "get_indices_arr" (get_ml_prog_state())) - [bv; fmlv; indsv; sv; vimapv] + [fmlv; indsv; sv; vimapv] (ARRAY fmlv fmllsv) (POSTv v. ARRAY fmlv fmllsv * &( - PAIR_TYPE (LIST_TYPE NUM) - (LIST_TYPE constraint_TYPE) - (get_indices b tcb fmlls inds s vimap) v)) + LIST_TYPE NUM + (get_indices fmlls inds s vimap) v)) Proof rw[]>> xcf "get_indices_arr" (get_ml_prog_state ())>> @@ -1934,9 +2021,10 @@ val check_red_arr = process_topdecs` case red_fast s idopt pfs of None => ( - let val bortcb = b orelse tcb in - case get_indices_arr bortcb fml inds s vimap of - (rinds,fmlls) => + let + val bortcb = b orelse tcb + val rinds = get_indices_arr fml inds s vimap in + case set_indices inds s vimap rinds of (inds',vimap') => let val nc = not_1 c val rsubs = do_rso ord s c obj @@ -1948,20 +2036,16 @@ val check_red_arr = process_topdecs` None => let val u = rollback_arr fml' id id' val goals = subst_indexes_arr s bortcb fml' rinds in - case red_cond_check fmlls nc pfs rsubs goals + case red_cond_check bortcb fml' inds' nc pfs rsubs goals of None => - (case set_indices inds s vimap rinds of - (inds',vimap') => - (fml', (inds', (vimap',id')))) + (fml', (inds', (vimap',id'))) | Some err => raise Fail (format_failure_2 lno ("Redundancy subproofs did not cover all subgoals. Info: " ^ err ^ ".") (print_subproofs_err rsubs goals)) end | Some cid => if check_contradiction_fml_arr b fml' cid then let val u = rollback_arr fml' id id' in - (case set_indices inds s vimap rinds of - (inds',vimap') => - (fml', (inds', (vimap',id')))) + (fml', (inds', (vimap',id'))) end else raise Fail (format_failure lno ("did not derive contradiction from index:" ^ Int.toString cid))) end @@ -1981,9 +2065,6 @@ Overload "ord_TYPE" = `` Overload "obj_TYPE" = `` OPTION_TYPE (PAIR_TYPE (LIST_TYPE (PAIR_TYPE INT NUM)) INT)`` -Overload "vimap_TYPE" = `` - SPTREE_SPT_TYPE (LIST_TYPE NUM)`` - Theorem check_red_arr_fast_spec: NUM lno lnov ∧ BOOL b bv ∧ @@ -2118,6 +2199,7 @@ Proof xvar>>xsimpl)>> pairarg_tac>>gs[]>> xlet_autop>> + xlet_autop>> fs[PAIR_TYPE_def]>> xmatch>> rpt xlet_autop>> @@ -2175,12 +2257,9 @@ Proof >- ( fs[red_cond_check_def]>> pairarg_tac>>fs[]>> - xlet_autop>> - fs[PAIR_TYPE_def]>> - xmatch>> rpt xlet_autop>> xcon>>xsimpl>> - simp[PAIR_TYPE_def]>> + fs[PAIR_TYPE_def]>> xsimpl) >- ( rpt xlet_autop>> @@ -2199,12 +2278,8 @@ Proof rw[]>>fs[]>> metis_tac[ARRAY_refl,NOT_EVERY,Fail_exn_def]) >> rpt xlet_autop>> - pairarg_tac>> - fs[PAIR_TYPE_def]>> - xmatch>> - rpt xlet_autop>> xcon>>xsimpl>> - simp[PAIR_TYPE_def]>> + fs[PAIR_TYPE_def]>> xsimpl QED @@ -2492,7 +2567,7 @@ QED Definition check_dom_list_def: check_dom_list spo obj fml inds id c s pfs idopt = - let (rinds,fmlls) = reindex F fml inds in + let rinds = reindex fml inds in let corels = core_fmlls fml rinds in let nc = not c in let fml_not_c = update_resize fml NONE (SOME (nc,F)) id in @@ -2506,7 +2581,7 @@ Definition check_dom_list_def: | SOME (fml',id') => let rfml = rollback fml' id id' in if do_dom_check idopt fml' rfml w - corels fmlls nc pfs dsubs then + corels rinds nc pfs dsubs then SOME (rfml,rinds,id') else NONE) End @@ -2529,9 +2604,8 @@ val res = translate map_snd_def; val check_dom_arr = process_topdecs` fun check_dom_arr lno spo obj fml inds id c s pfs idopt = - (case reindex_arr False fml inds of - (rinds,fmlls) => let + val rinds = reindex_arr fml inds val corels = core_fmlls_arr fml rinds val dsubs = do_dso spo s c obj val cpfs = extract_clauses_arr lno s False fml dsubs pfs [] @@ -2544,7 +2618,7 @@ val check_dom_arr = process_topdecs` None => let val u = rollback_arr fml' id id' val goals = core_subgoals s corels in - case red_cond_check fmlls nc pfs dsubs goals of + case red_cond_check False fml' rinds nc pfs dsubs goals of None => (fml',(rinds,id')) | Some err => raise Fail (format_failure_2 lno ("Dominance subproofs did not cover all subgoals. Info: " ^ err ^ ".") (print_subproofs_err dsubs goals)) @@ -2555,7 +2629,7 @@ val check_dom_arr = process_topdecs` (fml', (rinds,id')) end else raise Fail (format_failure lno ("did not derive contradiction from index:" ^ Int.toString cid))) - end)` |> append_prog; + end` |> append_prog; Theorem check_dom_arr_spec: NUM lno lnov ∧ @@ -2596,18 +2670,6 @@ Proof rw[]>> xcf "check_dom_arr" (get_ml_prog_state ())>> rw[]>> - xlet_autop>> - xlet`POSTv v. - ARRAY fmlv fmllsv * - &(PAIR_TYPE (LIST_TYPE NUM) - (LIST_TYPE constraint_TYPE) - (reindex F fmlls inds) v)` - >- ( - xapp>>xsimpl>> - EVAL_TAC)>> - Cases_on`reindex F fmlls inds`>> - fs[PAIR_TYPE_def]>> - xmatch>> rpt xlet_autop>> xlet`(POSTve (λv. @@ -2686,7 +2748,25 @@ Proof fs[OPTION_TYPE_def,do_dom_check_def,check_dom_list_def]>> xmatch >- ( - rpt xlet_autop>> + xlet_autop>> + xlet_autop>> + xlet`POSTv v. ARRAY fmlv' fmllsv'' * & BOOL F v` + >- + (xcon>>xsimpl)>> + qmatch_asmsub_abbrev_tac`LIST_REL _ fmlls'' fmllsv''`>> + xlet`POSTv resv. + ARRAY fmlv' fmllsv'' * + & ∃err. + OPTION_TYPE STRING_TYPE + (if + red_cond_check F fmlls'' (reindex fmlls inds) (not n) pfs + (dom_subgoals spo (subst_fun s) n obj) + (core_subgoals s (core_fmlls fmlls (reindex fmlls inds))) + then + NONE + else SOME err) resv` + >- + (xapp>>xsimpl)>> pop_assum mp_tac>>IF_CASES_TAC>> strip_tac>>fs[OPTION_TYPE_def]>>xmatch >- ( @@ -3171,7 +3251,7 @@ val check_cstep_arr = process_topdecs` (fml', (inds, (vimap, pc))) end) | Strengthentocore b => ( - let val inds' = fst (reindex_arr False fml inds) in + let val inds' = reindex_arr fml inds in if b then ( let val fml' = core_from_inds_arr lno fml inds' in @@ -3181,7 +3261,7 @@ val check_cstep_arr = process_topdecs` (fml,(inds', (vimap, set_tcb pc b))) end) | Loadorder nn xs => - let val inds' = fst (reindex_arr False fml inds) in + let val inds' = reindex_arr fml inds in case Alist.lookup (get_orders pc) nn of None => raise Fail (format_failure lno ("no such order: " ^ nn)) @@ -3311,7 +3391,6 @@ Proof rpt xlet_autop>> xraise>> xsimpl>> metis_tac[Fail_exn_def,ARRAY_refl])>> - pairarg_tac>>fs[]>> rpt xlet_autop>> xlet_auto >- (xsimpl>> simp (eq_lemmas()))>> @@ -3474,14 +3553,6 @@ Proof >- ( (* StrengthenToCore *) xmatch>> rpt xlet_autop>> - xlet`POSTv v. - ARRAY fmlv fmllsv * - &(PAIR_TYPE (LIST_TYPE NUM) (LIST_TYPE constraint_TYPE) - (reindex F fmlls inds) v)` - >- ( - xapp>>xsimpl>> - EVAL_TAC)>> - rpt xlet_autop>> xif >- ( xlet_autop>>xsimpl @@ -3499,14 +3570,6 @@ Proof >- ( (* LoadOrder*) xmatch>> rpt xlet_autop>> - xlet`POSTv v. - ARRAY fmlv fmllsv * - &(PAIR_TYPE (LIST_TYPE NUM) (LIST_TYPE constraint_TYPE) - (reindex F fmlls inds) v)` - >- ( - xapp>>xsimpl>> - EVAL_TAC)>> - rpt xlet_autop>> fs[get_orders_def]>> Cases_on`ALOOKUP pc.orders m`>>fs[OPTION_TYPE_def]>>xmatch >- ( From 09f24ecb64e7a83eb6bd05c4b8d6f27acea8bca6 Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Fri, 26 Jan 2024 23:31:08 +0800 Subject: [PATCH 14/38] fix npbc_array --- .../array/npbc_arrayProgScript.sml | 26 +++++++++---------- 1 file changed, 12 insertions(+), 14 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_arrayProgScript.sml b/examples/pseudo_bool/array/npbc_arrayProgScript.sml index 2495a82428..6460f78702 100644 --- a/examples/pseudo_bool/array/npbc_arrayProgScript.sml +++ b/examples/pseudo_bool/array/npbc_arrayProgScript.sml @@ -1707,9 +1707,6 @@ QED val hash_check = process_topdecs` fun hash_check fmlls proved lf = - case lf of - [] => None - | _ => let val hs = Array.array splim [] val u = mk_hashset_arr proved hs @@ -1744,14 +1741,6 @@ Theorem hash_check_spec: Proof rw[]>> xcf "hash_check" (get_ml_prog_state ())>> - Cases_on`ls` - >- ( - fs[LIST_TYPE_def]>> - xmatch>> - xcon>>xsimpl>> - EVAL_TAC)>> - fs[LIST_TYPE_def]>> - xmatch>> xlet_autop>> assume_tac (fetch "-" "splim_v_thm")>> xlet_auto>> @@ -1772,7 +1761,7 @@ Proof fs[LIST_REL_EL_EQN,Abbr`hs`])>> xapp>>xsimpl>> first_assum (irule_at Any)>> - qexists_tac`h::t`>>simp[LIST_TYPE_def]>> + qexists_tac`ls`>>simp[LIST_TYPE_def]>> unabbrev_all_tac>>fs[LIST_REL_EL_EQN] QED @@ -1829,6 +1818,8 @@ val red_cond_check = process_topdecs` case red_cond_check_pure extra pfs rsubs goals of None => Some "not all # subgoals present" | Some (x,ls) => + case ls of [] => None + | _ => let val fmlls = revalue_arr bortcb fml inds in hash_check fmlls x ls end` |> append_prog @@ -1867,14 +1858,21 @@ Proof xcon>>xsimpl)>> Cases_on`x`>>fs[PAIR_TYPE_def]>> xmatch>> + rename1` _ = SOME (proved,lf)`>> + Cases_on`lf`>> + fs[LIST_TYPE_def]>> + xmatch + >- ( + xcon>>xsimpl>> + EVAL_TAC)>> xlet_autop>> xapp>> xsimpl>> fs[]>> first_x_assum (irule_at Any)>> first_x_assum (irule_at Any)>> - first_x_assum (irule_at Any)>> - simp[] + qexists_tac`h::t`>> + simp[LIST_TYPE_def] QED (* Definition print_subgoal_def: From dbd74c50acc8ca53c2d0f291c8b476460f88322f Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sat, 27 Jan 2024 22:37:25 +0800 Subject: [PATCH 15/38] add # special case to npbc_check --- examples/pseudo_bool/npbc_checkScript.sml | 489 +++++++++++++++------- 1 file changed, 344 insertions(+), 145 deletions(-) diff --git a/examples/pseudo_bool/npbc_checkScript.sml b/examples/pseudo_bool/npbc_checkScript.sml index 2a5d9e2551..4c1101febc 100644 --- a/examples/pseudo_bool/npbc_checkScript.sml +++ b/examples/pseudo_bool/npbc_checkScript.sml @@ -726,9 +726,16 @@ Definition mem_constraint_def: if equal_constraint c x then T else mem_constraint c xs End +(* Checking triviality for a negated target c *) +Definition check_triv_def: + check_triv extra nc = + (check_contradiction (add extra nc) ∨ + check_contradiction nc) +End + (* Partition the formula goals into proved and non-proved For each non-proved goal, check if - 0) it is implied by extra (= ¬ C) + 0) it is implied by extra (= ¬ C) or trivial 1) it was already in the formula 2) it was already proved by another proofgoal (excluding #) *) @@ -740,7 +747,7 @@ Definition split_goals_def: PARTITION (λ(i,c). sptree$lookup i proved ≠ NONE) goals in let proved = MAP SND lp in let f = range fml in - EVERY (λ(i,c). c ∈ f ∨ imp extra c ∨ mem_constraint c proved) lf + EVERY (λ(i,c). c ∈ f ∨ check_triv extra (not c) ∨ mem_constraint c proved) lf End Triviality list_pair_eq_thm: @@ -776,6 +783,11 @@ Definition mk_subst_def: (mk_subst xs = INR (spt_to_vec (fromAList xs))) End +Definition check_hash_triv_def: + check_hash_triv extra ncs = + EXISTS (check_triv extra) ncs +End + (* The tcb flag indicates we're in to-core mode where it is guaranteed that the core formula implies derived *) Definition check_red_def: @@ -798,8 +810,11 @@ Definition check_red_def: let goals = toAList (map_opt (subst_opt w) gfml) in let (l,r) = extract_pids pfs LN LN in split_goals gfml nc l goals ∧ - EVERY (λid. lookup id r ≠ NONE) - (COUNT_LIST (LENGTH rsubs))) + EVERY (λ(id,cs). + lookup id r ≠ NONE ∨ + check_hash_triv nc cs + ) + (enumerate 0 rsubs)) | SOME cid => check_contradiction_fml b fml' cid) in if chk then @@ -1062,7 +1077,7 @@ QED Theorem split_goals_checked: split_goals fml e proved goals ∧ MEM (n,yy) goals ⇒ - yy ∈ range fml ∨ imp e yy ∨ + yy ∈ range fml ∨ check_triv e (not yy) ∨ ∃i. lookup i proved ≠ NONE ∧ MEM (i,yy) goals @@ -1164,6 +1179,53 @@ Proof gvs[AllCaseEqs()] QED +Theorem MEM_enumerate_index: + ∀k xs. + MEM (i,e) (enumerate k xs) ⇒ + ∃j. j < LENGTH xs ∧ i = k + j +Proof + Induct_on`xs`>>rw[miscTheory.enumerate_def] + >- intLib.ARITH_TAC>> + first_x_assum drule>> + rw[] + >- intLib.ARITH_TAC +QED + +Theorem MEM_enumerate_iff: + MEM ie (enumerate 0 xs) ⇔ + ∃i e. ie = (i,e) ∧ i < LENGTH xs ∧ EL i xs = e +Proof + Cases_on`ie`>> + rename1`MEM (i,e) _ `>> + Cases_on`i < LENGTH xs`>>fs[MEM_enumerate] + >- metis_tac[]>> + CCONTR_TAC>>fs[]>> + drule MEM_enumerate_index>> + rw[] +QED + +Theorem check_triv_unsatisfiable: + check_triv extra nc ⇒ + unsatisfiable (fml ∪ {extra} ∪ {nc}) +Proof + rw[check_triv_def]>> + drule check_contradiction_unsat>> + rw[unsatisfiable_def,satisfiable_def]>> + metis_tac[add_thm] +QED + +Theorem check_triv_unsatisfiable_2: + check_triv extra nc ∧ + extra ∈ fml ∧ nc ∈ fml + ⇒ + unsatisfiable fml +Proof + rw[check_triv_def]>> + drule check_contradiction_unsat>> + rw[unsatisfiable_def,satisfiable_def,satisfies_def]>> + metis_tac[add_thm] +QED + Theorem check_red_correct: id_ok fml id ∧ OPTION_ALL good_spo ord ∧ @@ -1224,55 +1286,89 @@ Proof \\ drule sat_implies_transitive \\ disch_then (fn th => DEP_REWRITE_TAC[th]) \\ simp [Once implies_explode] - \\ gvs[red_subgoals_def,MEM_COUNT_LIST,ADD1] + \\ gvs[red_subgoals_def,MEM_enumerate_iff,ADD1,AND_IMP_INTRO,PULL_EXISTS] \\ reverse (rw []) >- ( (* dominance *) rw[sat_implies_EL]>> - last_x_assum(qspec_then`SUC n` assume_tac)>> + last_x_assum(qspec_then`SUC n` mp_tac)>> gvs[]>> - drule_all lookup_extract_pids_r>> rw[] - \\ drule extract_clauses_MEM_INR - \\ disch_then drule - \\ fs[EL] - \\ DEP_REWRITE_TAC [EL_APPEND_EQN] \\ simp[] - \\ rw[] - \\ first_x_assum drule \\ strip_tac - \\ gs[EL_MAP] - \\ drule unsatisfiable_not_sat_implies - \\ simp[range_insert] - \\ metis_tac[INSERT_SING_UNION,UNION_COMM]) + PURE_REWRITE_TAC[METIS_PROVE [] ``((P ⇒ Q) ⇒ R) ⇔ (~P ∨ Q) ⇒ R``]>> + strip_tac + >- ( + drule_all lookup_extract_pids_r>> rw[] + \\ drule extract_clauses_MEM_INR + \\ disch_then drule + \\ fs[EL] + \\ DEP_REWRITE_TAC [EL_APPEND_EQN] \\ simp[] + \\ rw[] + \\ first_x_assum drule \\ strip_tac + \\ gs[EL_MAP] + \\ drule unsatisfiable_not_sat_implies + \\ simp[range_insert] + \\ metis_tac[INSERT_SING_UNION,UNION_COMM]) + >- ( + fs[check_hash_triv_def] + \\ pop_assum mp_tac + \\ DEP_REWRITE_TAC [EL_APPEND_EQN] + \\ simp[EL_MAP] + \\ strip_tac + \\ match_mp_tac unsatisfiable_not_sat_implies + \\ metis_tac[check_triv_unsatisfiable]) ) >- ( (* objective *) - Cases_on`obj`>> - last_x_assum(qspec_then`SUC(LENGTH (dom_subst (subst_fun (mk_subst s)) ord))` assume_tac)>> - fs[]>> - drule_all lookup_extract_pids_r>>rw[] - \\ drule extract_clauses_MEM_INR - \\ disch_then drule - \\ fs[EL] - \\ DEP_REWRITE_TAC [EL_APPEND2] - \\ simp[] - \\ rw[] - \\ first_x_assum drule \\ strip_tac - \\ gs[] - \\ drule unsatisfiable_not_sat_implies - \\ simp[range_insert] - \\ metis_tac[INSERT_SING_UNION,UNION_COMM]) + Cases_on`obj`>> gvs[]>> + last_x_assum(qspec_then`SUC(LENGTH (dom_subst (subst_fun (mk_subst s)) ord))` mp_tac)>> + gvs[]>> + PURE_REWRITE_TAC[METIS_PROVE [] ``((P ⇒ Q) ⇒ R) ⇔ (~P ∨ Q) ⇒ R``]>> + strip_tac + >- ( + drule_all lookup_extract_pids_r>>rw[] + \\ drule extract_clauses_MEM_INR + \\ disch_then drule + \\ fs[EL] + \\ DEP_REWRITE_TAC [EL_APPEND2] + \\ simp[] + \\ rw[] + \\ first_x_assum drule \\ strip_tac + \\ gs[] + \\ drule unsatisfiable_not_sat_implies + \\ simp[range_insert] + \\ metis_tac[INSERT_SING_UNION,UNION_COMM]) + >- ( + fs[check_hash_triv_def] + \\ pop_assum mp_tac + \\ DEP_REWRITE_TAC [EL_APPEND_EQN] + \\ simp[EL_MAP] + \\ strip_tac + \\ match_mp_tac unsatisfiable_not_sat_implies + \\ metis_tac[check_triv_unsatisfiable]) + ) >- ( (* redundancy #0 *) - last_x_assum(qspec_then`0` assume_tac)>> - fs[]>> - drule_all lookup_extract_pids_r>>rw[] - \\ drule extract_clauses_MEM_INR - \\ disch_then drule - \\ fs[] - \\ rw[] - \\ first_x_assum drule \\ strip_tac - \\ fs[] - \\ drule unsatisfiable_not_sat_implies - \\ simp[range_insert] - \\ metis_tac[INSERT_SING_UNION,UNION_COMM]) + last_x_assum(qspec_then`0` mp_tac)>> + gvs[]>> + PURE_REWRITE_TAC[METIS_PROVE [] ``((P ⇒ Q) ⇒ R) ⇔ (~P ∨ Q) ⇒ R``]>> + strip_tac + >- ( + drule_all lookup_extract_pids_r>>rw[] + \\ drule extract_clauses_MEM_INR + \\ disch_then drule + \\ fs[] + \\ rw[] + \\ first_x_assum drule \\ strip_tac + \\ fs[] + \\ drule unsatisfiable_not_sat_implies + \\ simp[range_insert] + \\ metis_tac[INSERT_SING_UNION,UNION_COMM]) + >- ( + fs[check_hash_triv_def] + \\ pop_assum mp_tac + \\ DEP_REWRITE_TAC [EL_APPEND_EQN] + \\ simp[EL_MAP] + \\ strip_tac + \\ match_mp_tac unsatisfiable_not_sat_implies + \\ metis_tac[check_triv_unsatisfiable])) (* rest of redundancy *) \\ gvs [GSYM unsat_iff_implies] \\ Cases_on ‘subst_opt (subst_fun (mk_subst s)) x'’ \\ fs [] @@ -1297,7 +1393,8 @@ Proof simp[]>> metis_tac[range_mk_core_fml,in_core_only_fml_or_left]) >- ( - fs[satisfiable_def,not_thm,satisfies_def]>> + drule check_triv_unsatisfiable>> + fs[unsatisfiable_def,satisfiable_def,not_thm,satisfies_def]>> drule subst_opt_SOME >> metis_tac[not_thm,imp_thm,in_core_only_fml_or_left]) \\ drule_all lookup_extract_pids_l>>rw[] @@ -1770,7 +1867,11 @@ Definition check_reflexivity_def: (case check_subproofs cpfs F fml id of SOME (fml',id') => let (l,r) = extract_pids pfs LN LN in - EVERY (λid. lookup id r ≠ NONE) (COUNT_LIST (LENGTH dsubs)) + EVERY (λ(id,cs). + lookup id r ≠ NONE ∨ + EXISTS check_contradiction cs + ) + (enumerate 0 dsubs) | _ => F) End @@ -1786,7 +1887,12 @@ Definition check_transitivity_def: (case check_subproofs cpfs F fml id of SOME (fml',id') => let (l,r) = extract_pids pfs LN LN in - if EVERY (λid. lookup id r ≠ NONE) (COUNT_LIST (LENGTH dsubs)) + if + EVERY (λ(id,cs). + lookup id r ≠ NONE ∨ + EXISTS check_contradiction cs + ) + (enumerate 0 dsubs) then SOME id' else NONE | _ => NONE) @@ -1931,8 +2037,12 @@ Definition check_change_obj_def: NONE => NONE | SOME (fml',id') => let (l,r) = extract_pids pfs LN LN in - if EVERY (λid. lookup id r ≠ NONE) - (COUNT_LIST (LENGTH csubs)) + if + EVERY (λ(id,cs). + lookup id r ≠ NONE ∨ + EXISTS check_contradiction cs + ) + (enumerate 0 csubs) then SOME (fc',id') else NONE)) @@ -2014,8 +2124,11 @@ Definition check_cstep_def: let (l,r) = extract_pids pfs LN LN in let gfml = mk_core_fml F fml in split_goals gfml nc l goals ∧ - EVERY (λid. lookup id r ≠ NONE) - (COUNT_LIST (LENGTH dsubs)) + EVERY (λ(id,cs). + lookup id r ≠ NONE ∨ + check_hash_triv nc cs + ) + (enumerate 0 dsubs) | SOME cid => check_contradiction_fml F fml' cid) in if check then @@ -2400,7 +2513,7 @@ Proof `pc.id ∉ domain fml` by fs[id_ok_def]>> fs[core_only_fml_F_insert_b]>> fs[EVERY_MEM,MEM_MAP,EXISTS_PROD]>> - gvs[dom_subgoals_def,MEM_COUNT_LIST,ADD1]>> + gvs[dom_subgoals_def,MEM_enumerate_iff,ADD1,AND_IMP_INTRO,PULL_EXISTS]>> CONJ_TAC >- ( (* core constraints*) rw[sat_implies_def,satisfies_def] @@ -2423,10 +2536,13 @@ Proof drule subst_opt_SOME >> rw[]>> metis_tac[not_thm]) >- ( - fs[satisfiable_def,not_thm,satisfies_def]>> + drule check_triv_unsatisfiable>> + disch_then(qspec_then`{}` mp_tac)>> + fs[unsatisfiable_def,satisfiable_def,not_thm,satisfies_def]>> drule subst_opt_SOME >> fs[range_def]>> - rw[]>> metis_tac[not_thm,imp_thm]) + rw[]>> + metis_tac[not_thm,imp_thm]) \\ drule_all lookup_extract_pids_l>>rw[] \\ drule_all extract_clauses_MEM_INL \\ strip_tac @@ -2453,29 +2569,75 @@ Proof fs[core_only_fml_def]>> simp[GSYM LIST_TO_SET_MAP]>> rw[sat_implies_EL,EL_MAP]>> - last_x_assum(qspec_then`n` assume_tac)>> + last_x_assum(qspec_then`n` mp_tac)>> gvs[dom_subst_def]>> - drule_all lookup_extract_pids_r>> - rw[]>> - drule extract_clauses_MEM_INR>> - disch_then drule>> - fs[EL_MAP]>> - DEP_REWRITE_TAC [EL_APPEND_EQN] >> simp[]>> - simp[EL_MAP]>> - rw[]>> - first_x_assum drule >> strip_tac>> - gs[]>> - drule unsatisfiable_not_sat_implies>> - simp[lookup_list_list_insert,ALOOKUP_ZIP_MAP]>> - simp[range_insert]>> - metis_tac[INSERT_SING_UNION,UNION_COMM])>> + PURE_REWRITE_TAC[METIS_PROVE [] ``((P ⇒ Q) ⇒ R) ⇔ (~P ∨ Q) ⇒ R``]>> + strip_tac + >- ( + drule_all lookup_extract_pids_r>> + rw[]>> + drule extract_clauses_MEM_INR>> + disch_then drule>> + fs[EL_MAP]>> + DEP_REWRITE_TAC [EL_APPEND_EQN] >> simp[]>> + simp[EL_MAP]>> + rw[]>> + first_x_assum drule >> strip_tac>> + gs[]>> + drule unsatisfiable_not_sat_implies>> + simp[lookup_list_list_insert,ALOOKUP_ZIP_MAP]>> + simp[range_insert]>> + metis_tac[INSERT_SING_UNION,UNION_COMM]) + >- ( + fs[check_hash_triv_def] + \\ pop_assum mp_tac + \\ DEP_REWRITE_TAC [EL_APPEND_EQN] + \\ simp[EL_MAP] + \\ simp[lookup_list_list_insert,ALOOKUP_ZIP_MAP] + \\ strip_tac + \\ match_mp_tac unsatisfiable_not_sat_implies + \\ metis_tac[check_triv_unsatisfiable]) + )>> CONJ_TAC >- ( (* negated order constraint *) fs[core_only_fml_def]>> - last_x_assum(qspec_then`LENGTH (dom_subst (subst_fun (mk_subst l)) (SOME ((x0,x1,x2),x3)))` assume_tac)>> + last_x_assum(qspec_then`LENGTH (dom_subst (subst_fun (mk_subst l)) (SOME ((x0,x1,x2),x3)))` mp_tac)>> gs[ADD1]>> + PURE_REWRITE_TAC[METIS_PROVE [] ``((P ⇒ Q) ⇒ R) ⇔ (~P ∨ Q) ⇒ R``]>> + strip_tac + >- ( + drule_all lookup_extract_pids_r>> + simp[]>> rw[] + \\ drule extract_clauses_MEM_INR + \\ disch_then drule + \\ fs[] + \\ DEP_REWRITE_TAC [EL_APPEND2] + \\ simp[] + \\ rw[] + \\ first_x_assum drule \\ strip_tac + \\ gs[neg_dom_subst_def,lookup_list_list_insert,ALOOKUP_ZIP_MAP,range_insert] + \\ metis_tac[INSERT_SING_UNION,UNION_COMM,LIST_TO_SET_MAP]) + >- ( + fs[check_hash_triv_def] + \\ pop_assum mp_tac + \\ DEP_REWRITE_TAC [EL_APPEND_EQN] + \\ gs[neg_dom_subst_def,lookup_list_list_insert,ALOOKUP_ZIP_MAP,range_insert,EXISTS_MEM,MEM_MAP] + \\ strip_tac \\ rw[] + \\ drule check_triv_unsatisfiable_2 + \\ disch_then match_mp_tac + \\ simp[]) + )>> + (* objective constraint *) + fs[core_only_fml_def]>> + Cases_on`pc.obj`>> + simp[]>> + last_x_assum(qspec_then`SUC(LENGTH (dom_subst (subst_fun (mk_subst l)) (SOME ((x0,x1,x2),x3))))` mp_tac)>> + gs[ADD1]>> + PURE_REWRITE_TAC[METIS_PROVE [] ``((P ⇒ Q) ⇒ R) ⇔ (~P ∨ Q) ⇒ R``]>> + strip_tac + >- ( drule_all lookup_extract_pids_r>> - simp[]>> rw[] + simp[]>>rw[] \\ drule extract_clauses_MEM_INR \\ disch_then drule \\ fs[] @@ -2483,26 +2645,19 @@ Proof \\ simp[] \\ rw[] \\ first_x_assum drule \\ strip_tac - \\ gs[neg_dom_subst_def,lookup_list_list_insert,ALOOKUP_ZIP_MAP,range_insert] - \\ metis_tac[INSERT_SING_UNION,UNION_COMM,LIST_TO_SET_MAP])>> - (* objective constraint *) - fs[core_only_fml_def]>> - Cases_on`pc.obj`>> - simp[]>> - last_x_assum(qspec_then`SUC(LENGTH (dom_subst (subst_fun (mk_subst l)) (SOME ((x0,x1,x2),x3))))` assume_tac)>> - gs[ADD1]>> - drule_all lookup_extract_pids_r>> - simp[]>>rw[] - \\ drule extract_clauses_MEM_INR - \\ disch_then drule - \\ fs[] - \\ DEP_REWRITE_TAC [EL_APPEND2] - \\ simp[] - \\ rw[] - \\ first_x_assum drule \\ strip_tac - \\ gs[range_insert] - \\ drule unsatisfiable_not_sat_implies - \\ metis_tac[INSERT_SING_UNION,UNION_COMM])>> + \\ gs[range_insert] + \\ drule unsatisfiable_not_sat_implies + \\ metis_tac[INSERT_SING_UNION,UNION_COMM]) + >- ( + fs[check_hash_triv_def] + \\ pop_assum mp_tac + \\ DEP_REWRITE_TAC [EL_APPEND_EQN] + \\ simp[EL_MAP] + \\ simp[lookup_list_list_insert,ALOOKUP_ZIP_MAP] + \\ strip_tac + \\ match_mp_tac unsatisfiable_not_sat_implies + \\ metis_tac[check_triv_unsatisfiable]) + )>> CONJ_TAC >- ( pop_assum mp_tac>> DEP_REWRITE_TAC[core_only_fml_F_insert_b]>> @@ -2724,22 +2879,33 @@ Proof rw[]>> simp[GSYM LIST_TO_SET_MAP]>> rw[sat_implies_EL]>> - fs[EVERY_MEM,MEM_COUNT_LIST]>> - first_x_assum drule>> - rw[]>> - drule_all lookup_extract_pids_r>> - simp[]>> - rw[]>> - drule extract_clauses_MEM_INR>> - disch_then drule>> - simp[EL_MAP]>> - rw[]>> + fs[EVERY_MEM,MEM_enumerate_iff,PULL_EXISTS]>> first_x_assum drule>> - simp[]>>rw[]>> - drule unsatisfiable_not_sat_implies>> - simp[core_only_fml_def,Abbr`fmll`] >> - once_rewrite_tac [subst_eta] >> fs [] >> - gvs [ALOOKUP_APPEND,ALOOKUP_ZIP_MAP])>> + PURE_REWRITE_TAC[METIS_PROVE [] ``((P ⇒ Q) ⇒ R) ⇔ (~P ∨ Q) ⇒ R``]>> + strip_tac + >- ( + drule_all lookup_extract_pids_r>> + simp[]>> + rw[]>> + drule extract_clauses_MEM_INR>> + disch_then drule>> + simp[EL_MAP]>> + rw[]>> + first_x_assum drule>> + simp[]>>rw[]>> + drule unsatisfiable_not_sat_implies>> + simp[core_only_fml_def,Abbr`fmll`] >> + once_rewrite_tac [subst_eta] >> fs [] >> + gvs [ALOOKUP_APPEND,ALOOKUP_ZIP_MAP]) + >- ( + gvs[EL_MAP]>> + match_mp_tac unsatisfiable_not_sat_implies>> + simp[]>> + drule check_contradiction_unsat>> + simp[unsatisfiable_def,satisfiable_def]>> + once_rewrite_tac [subst_eta] >> fs [] >> + rw[ALOOKUP_ZIP_MAP] ) + )>> PairCases_on`p`>> match_mp_tac (transitive_po_of_spo |> SIMP_RULE std_ss [AND_IMP_INTRO] |> GEN_ALL)>> gvs[check_transitivity_def,check_ws_def]>> @@ -2761,31 +2927,43 @@ Proof rw[]>> simp[GSYM LIST_TO_SET_MAP]>> rw[sat_implies_EL]>> - fs[MEM_COUNT_LIST]>> - pairarg_tac>>gs[]>> - first_x_assum drule>> - rw[]>> - drule_all lookup_extract_pids_r>> - simp[]>> - rw[]>> - drule extract_clauses_MEM_INR>> - disch_then drule>> - simp[EL_MAP]>> - rw[]>> - fs[EVERY_MEM]>> + fs[MEM_enumerate_iff]>> + pairarg_tac>>gs[PULL_EXISTS]>> first_x_assum drule>> - simp[]>>rw[]>> - drule unsatisfiable_not_sat_implies>> - simp[core_only_fml_build_fml,Abbr`fmll`] >> - once_rewrite_tac [subst_eta] >> fs [] >> - gvs [ALOOKUP_APPEND,ALOOKUP_ZIP_MAP] >> - qmatch_goalsub_abbrev_tac ‘subst f1’ >> strip_tac >> - qmatch_goalsub_abbrev_tac ‘subst f2’ >> - qsuff_tac ‘f1 = f2’ >- - (rw [] >> gvs []) >> - unabbrev_all_tac >> - fs [FUN_EQ_THM] >> rw [] >> - ntac 4 (CASE_TAC >> fs []) ) + PURE_REWRITE_TAC[METIS_PROVE [] ``((P ⇒ Q) ⇒ R) ⇔ (~P ∨ Q) ⇒ R``]>> + strip_tac + >- ( + rw[]>> + drule_all lookup_extract_pids_r>> + simp[]>> + rw[]>> + drule extract_clauses_MEM_INR>> + disch_then drule>> + simp[EL_MAP]>> + rw[]>> + fs[EVERY_MEM]>> + first_x_assum drule>> + simp[]>>rw[]>> + drule unsatisfiable_not_sat_implies>> + simp[core_only_fml_build_fml,Abbr`fmll`] >> + once_rewrite_tac [subst_eta] >> fs [] >> + gvs [ALOOKUP_APPEND,ALOOKUP_ZIP_MAP] >> + qmatch_goalsub_abbrev_tac ‘subst f1’ >> strip_tac >> + qmatch_goalsub_abbrev_tac ‘subst f2’ >> + qsuff_tac ‘f1 = f2’ >- + (rw [] >> gvs []) >> + unabbrev_all_tac >> + fs [FUN_EQ_THM] >> rw [] >> + ntac 4 (CASE_TAC >> fs []) ) + >- ( + gvs[EL_MAP]>> + match_mp_tac unsatisfiable_not_sat_implies>> + simp[]>> + drule check_contradiction_unsat>> + simp[unsatisfiable_def,satisfiable_def]>> + once_rewrite_tac [subst_eta] >> fs [] >> + rw[ALOOKUP_ZIP_MAP] ) + ) >- ( (* Obj *) strip_tac>> @@ -2855,20 +3033,40 @@ Proof disch_then(qspecl_then[`pfs`,`T`] mp_tac)>>simp[]>> strip_tac>> fs[EVERY_MEM,MEM_MAP,EXISTS_PROD]>> - gvs[change_obj_subgoals_def,MEM_COUNT_LIST,ADD1]>> - first_assum(qspec_then`0` mp_tac)>> - first_x_assum(qspec_then`1` mp_tac)>> - simp[]>> - strip_tac>> drule_all lookup_extract_pids_r>> - simp[]>> strip_tac>> - drule_all extract_clauses_MEM_INR>> - simp[]>> strip_tac>> - strip_tac>> drule_all lookup_extract_pids_r>> - simp[]>> strip_tac>> - drule_all extract_clauses_MEM_INR>> - simp[]>> strip_tac>> - res_tac >>fs[]>> + gvs[change_obj_subgoals_def,MEM_enumerate_iff,ADD1,PULL_EXISTS]>> qmatch_asmsub_rename_tac`model_bounding fnew fold`>> + `unsatisfiable (core_only_fml T fml ∪ {not (model_bounding fnew fold)})` by ( + first_x_assum(qspec_then`0` mp_tac)>> + simp[]>> + PURE_REWRITE_TAC[METIS_PROVE [] ``((P ⇒ Q) ⇒ R) ⇔ (~P ∨ Q) ⇒ R``]>> + strip_tac + >- ( + drule_all lookup_extract_pids_r>> + simp[]>> strip_tac>> + drule_all extract_clauses_MEM_INR>> + simp[]>> + strip_tac>> + first_x_assum drule>>simp[]) + >- ( + drule check_contradiction_unsat>> + simp[unsatisfiable_def,satisfiable_def] + ))>> + `unsatisfiable (core_only_fml T fml ∪ {not (model_bounding fold fnew)})` by ( + first_x_assum(qspec_then`1` mp_tac)>> + simp[]>> + PURE_REWRITE_TAC[METIS_PROVE [] ``((P ⇒ Q) ⇒ R) ⇔ (~P ∨ Q) ⇒ R``]>> + strip_tac + >- ( + drule_all lookup_extract_pids_r>> + simp[]>> strip_tac>> + drule_all extract_clauses_MEM_INR>> + simp[]>> + strip_tac>> + first_x_assum drule>>simp[]) + >- ( + drule check_contradiction_unsat>> + simp[unsatisfiable_def,satisfiable_def] + ))>> `∀w. satisfies w (core_only_fml T fml) ⇒ eval_obj (SOME fold) w = @@ -2899,7 +3097,8 @@ Proof metis_tac[])>> fs[valid_conf_def]>> rw[bimp_obj_def,imp_obj_def,sat_obj_le_def]>> - asm_exists_tac>>simp[]) + first_x_assum drule>> + metis_tac[]) >- (every_case_tac>>rw[]) QED From 612df4ef09747dadd3bb3cad638149a1a6020183 Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sun, 28 Jan 2024 00:19:18 +0800 Subject: [PATCH 16/38] propagate through _array --- .../array/npbc_arrayProgScript.sml | 48 ++++++++++--- .../pseudo_bool/array/npbc_listScript.sml | 68 +++++++++---------- 2 files changed, 69 insertions(+), 47 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_arrayProgScript.sml b/examples/pseudo_bool/array/npbc_arrayProgScript.sml index 6460f78702..2ad728f1a4 100644 --- a/examples/pseudo_bool/array/npbc_arrayProgScript.sml +++ b/examples/pseudo_bool/array/npbc_arrayProgScript.sml @@ -1772,7 +1772,24 @@ Definition red_cond_check_def: let (l,r) = extract_pids pfs LN LN in let fmlls = revalue bortcb fml inds in split_goals_hash fmlls extra l goals ∧ - EVERY (λid. lookup id r ≠ NONE) (COUNT_LIST (LENGTH rsubs)) + EVERY (λ(id,cs). + lookup id r ≠ NONE ∨ + check_hash_triv extra cs + ) + (enumerate 0 rsubs) +End + +Definition lookup_hash_triv_def: + lookup_hash_triv r extra (id,cs) = + case sptree$lookup id r of + NONE => check_hash_triv extra cs + | SOME _ => T +End + +Definition every_check_hash_triv_def: + every_check_hash_triv r extra rsubs = + EVERY (lookup_hash_triv r extra) + (enumerate 0 rsubs) End Definition red_cond_check_pure_def: @@ -1781,16 +1798,25 @@ Definition red_cond_check_pure_def: (rsubs:((int # num) list # num) list list) (goals:(num # (int # num) list # num) list) = let (l,r) = extract_pids pfs LN LN in if - EVERY (λid. lookup id r ≠ NONE) (COUNT_LIST (LENGTH rsubs)) + every_check_hash_triv r extra rsubs then let (lp,lf) = PARTITION (λ(i,c). lookup i l ≠ NONE) goals in - let lf = FILTER (λc. ¬(imp extra c)) (MAP SND lf) in + let lf = FILTER (λc. ¬check_triv extra (not c)) (MAP SND lf) in let proved = MAP SND lp in SOME (proved,lf) else NONE End +Theorem lookup_hash_triv_fun_eq: + lookup_hash_triv r extra = + (λ(id,cs). lookup id r = NONE ⇒ check_hash_triv extra cs) +Proof + rw[FUN_EQ_THM]>> + pairarg_tac>>fs[lookup_hash_triv_def]>> + every_case_tac>>gvs[] +QED + Theorem red_cond_check_eq: red_cond_check bortcb fml inds extra pfs rsubs goals = case red_cond_check_pure extra pfs rsubs goals of @@ -1800,17 +1826,19 @@ Theorem red_cond_check_eq: let hs = mk_hashset fmlls (mk_hashset x (REPLICATE splim [])) in EVERY (λc. in_hashset c hs) ls Proof - rw[red_cond_check_def,red_cond_check_pure_def]>> + rw[red_cond_check_def,red_cond_check_pure_def,every_check_hash_triv_def]>> pairarg_tac>>fs[split_goals_hash_def]>> - IF_CASES_TAC>>fs[]>> - rpt (pairarg_tac>>fs[]) + rpt (pairarg_tac>>fs[])>> + rw[lookup_hash_triv_fun_eq] QED -val res = translate COUNT_LIST_AUX_def; -val res = translate COUNT_LIST_compute; +val res = translate npbc_checkTheory.check_triv_def; +val res = translate npbc_checkTheory.check_hash_triv_def; +val res = translate miscTheory.enumerate_def; val res = translate PART_DEF; val res = translate PARTITION_DEF; - +val res = translate lookup_hash_triv_def; +val res = translate every_check_hash_triv_def; val res = translate red_cond_check_pure_def; val red_cond_check = process_topdecs` @@ -2951,7 +2979,7 @@ val check_change_obj_arr = process_topdecs ` case check_subproofs_arr lno cpfs b fml id id of (fml', id') => let val u = rollback_arr fml' id id' in - if do_change_obj_check pfs then + if do_change_obj_check pfs csubs then (fml',(fc',id')) else raise Fail (format_failure lno ("Objective change subproofs did not cover all subgoals. Expected: #[1-2]")) end diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 2fb6b89ff3..125013c1bc 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -718,7 +718,7 @@ Definition split_goals_hash_def: (goals:(num # (int # num) list # num) list) = let (lp,lf) = PARTITION (λ(i,c). lookup i proved ≠ NONE) goals in - let lf = FILTER (λc. ¬(imp extra c)) (MAP SND lf) in + let lf = FILTER (λc. ¬check_triv extra (not c)) (MAP SND lf) in let proved = MAP SND lp in let hs = mk_hashset fmlls (mk_hashset proved (REPLICATE splim [])) in EVERY (λc. in_hashset c hs) lf @@ -733,7 +733,11 @@ Definition do_red_check_def: let (l,r) = extract_pids pfs LN LN in let fmlls = revalue (b ∨ tcb) rfml inds in split_goals_hash fmlls extra l goals ∧ - EVERY (λid. lookup id r ≠ NONE) (COUNT_LIST (LENGTH rsubs)) + EVERY (λ(id,cs). + lookup id r ≠ NONE ∨ + check_hash_triv extra cs + ) + (enumerate 0 rsubs) | SOME cid => check_contradiction_fml_list b fml cid End @@ -1157,28 +1161,17 @@ Theorem revalue_aux_SUBSET: set vacc ⊆ core_only_fml b fml ⇒ set (revalue_aux b fmlls inds vacc) ⊆ core_only_fml b fml Proof - cheat -QED - -(* -SND_reindex_characterize: - fml_rel fml fmlls ∧ - reindex b fmlls inds = (is,vs) ⇒ - set vs ⊆ core_only_fml b fml -Proof - rw[reindex_def,reindex_aux]>> - simp[SUBSET_DEF,MEM_MAP,MEM_FILTER,PULL_EXISTS,range_def]>> - rw[]>> - fs[IS_SOME_EXISTS]>> + Induct>>rw[revalue_aux_def]>> + every_case_tac>>rw[]>> + first_x_assum match_mp_tac>>fs[]>> drule fml_rel_lookup_core_only>> - rw[]>> - gvs[lookup_core_only_def,core_only_fml_def,AllCaseEqs()] - >- - metis_tac[]>> - rw[]>>fs[]>> + simp[lookup_core_only_list_def]>> + disch_then(qspecl_then[`h`,`b`] mp_tac)>> + simp[]>>rw[]>>gvs[]>> + gvs[core_only_fml_def,lookup_core_only_def,AllCaseEqs()]>> + rw[]>>gvs[]>> metis_tac[] QED -*) Theorem revalue_SUBSET: fml_rel fml fmlls ==> @@ -1365,9 +1358,7 @@ Proof fs[EVERY_FILTER,EVERY_MAP]>> qpat_x_assum`EVERY _ _`mp_tac>> match_mp_tac MONO_EVERY>> simp[FORALL_PROD, METIS_PROVE []``(¬P ⇒ Q) ⇔ P ∨ Q``]>> - rw[] - >- - simp[]>> + rw[]>>simp[]>> drule in_hashset_mk_hashset>> rw[] >- fs[MEM_MAP,SUBSET_DEF]>> @@ -1921,7 +1912,11 @@ Definition do_dom_check_def: let goals = MAP_OPT (subst_opt w) indcore in let (l,r) = extract_pids pfs LN LN in - if EVERY (λid. lookup id r ≠ NONE) (COUNT_LIST (LENGTH dsubs)) + if + EVERY (λ(id,cs). + lookup id r ≠ NONE ∨ + check_hash_triv extra cs) + (enumerate 0 dsubs) then let fmlls = revalue F rfml rinds in split_goals_hash fmlls extra l goals @@ -1985,17 +1980,20 @@ Definition all_core_list_def: else NONE) End -Definition do_change_obj_check_def: - do_change_obj_check pfs = - let (l,r) = extract_pids pfs LN LN in - lookup 0 r ≠ NONE ∧ - lookup 1 r ≠ NONE -End - Definition emp_vec_def: emp_vec = INR (Vector []) End +Definition do_change_obj_check_def: + do_change_obj_check pfs csubs = + let (l,r) = extract_pids pfs LN LN in + EVERY (λ(id,cs). + lookup id r ≠ NONE ∨ + EXISTS check_contradiction cs + ) + (enumerate 0 csubs) +End + Definition check_change_obj_list_def: check_change_obj_list b fml id obj fc' pfs ⇔ case obj of NONE => NONE @@ -2009,7 +2007,7 @@ Definition check_change_obj_list_def: NONE => NONE | SOME (fml',id') => let rfml = rollback fml' id id' in - if do_change_obj_check pfs then + if do_change_obj_check pfs csubs then SOME (rfml,fc',id') else NONE) End @@ -2472,10 +2470,6 @@ Proof fs[do_change_obj_check_def]>> pairarg_tac>>fs[]>> strip_tac>>simp[]>> - CONJ_TAC>- ( - `COUNT_LIST (LENGTH (change_obj_subgoals fc (mk_diff_obj b fc p) )) = [0;1]` by - EVAL_TAC>> - simp[])>> drule check_subproofs_list_id>> drule check_subproofs_list_id_upper>> drule check_subproofs_list_mindel>> From 98f9c59d0c401207b239bb05f101fc9cd4e66b2e Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sun, 28 Jan 2024 00:52:42 +0800 Subject: [PATCH 17/38] remove some cheats in npbc_list --- .../pseudo_bool/array/npbc_listScript.sml | 47 ++++++++++++++----- 1 file changed, 34 insertions(+), 13 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 125013c1bc..3d8d3b5071 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -1462,8 +1462,14 @@ Theorem vimap_rel_get_indices_set_indices: vimap_rel fmlls vimap' Proof rw[get_indices_def,set_indices_def] >> - gvs[AllCaseEqs()]>> - cheat + gvs[AllCaseEqs(),vimap_rel_def]>> + every_case_tac>>rw[lookup_insert]>> + first_x_assum (drule_at Any)>>rw[] + >- + (CCONTR_TAC>>gvs[]>> first_x_assum drule>>rw[])>> + first_x_assum drule_all>> + rw[reindex_characterize,MEM_FILTER]>> + simp[any_el_ALT] QED Theorem fml_rel_check_red_list: @@ -1559,23 +1565,22 @@ Proof metis_tac[])*))>> drule subst_indexes_MEM>> rw[MEM_toAList,lookup_map_opt]>> - cheat - (* - drule_all FST_reindex_partial_characterize>> - TOP_CASE_TAC>>rw[]>>gvs[]>> + gvs[reindex_characterize]>> fs[rollback_def,lookup_core_only_list_list_delete_list,MEM_MAP,MEM_COUNT_LIST,MEM_FILTER]>> - `p_1 < id` by ( + `p_1 < id'` by ( CCONTR_TAC>>fs[]>> - last_x_assum(qspec_then`p_1` mp_tac)>> + first_x_assum(qspec_then`p_1` mp_tac)>> impl_tac>-fs[]>> - metis_tac[option_CLAUSES])>> + gvs[lookup_core_only_list_def]>> + CCONTR_TAC>>gvs[])>> + `p_1 < id` by intLib.ARITH_TAC>> simp[lookup_mk_core_fml]>> `lookup_core_only (b ∨ tcb) fml p_1 = SOME c'` by ( qpat_x_assum`fml_rel fml _` assume_tac>> drule (GSYM fml_rel_lookup_core_only)>> strip_tac>>fs[]>> gvs[lookup_core_only_list_def,AllCaseEqs()])>> - fs[]*))>> + fs[])>> match_mp_tac (GEN_ALL fml_rel_check_contradiction_fml)>> metis_tac[])>> CONJ_ASM1_TAC>- ( @@ -2265,6 +2270,22 @@ Proof fs[transitive_def] QED +Theorem vimap_rel_core_from_inds: + ∀l fmlls fmlls'. + vimap_rel fmlls vimap ∧ + core_from_inds fmlls l = SOME fmlls' ⇒ + vimap_rel fmlls' vimap +Proof + Induct>>rw[core_from_inds_def]>> + gvs[AllCaseEqs()]>> + first_x_assum match_mp_tac>> + first_x_assum (irule_at Any)>> + fs[vimap_rel_def]>>rw[]>> + gvs[update_resize_def]>>every_case_tac>> + gvs[EL_LUPDATE,EL_APPEND_EQN]>>every_case_tac>> + gvs[EL_REPLICATE,any_el_ALT] +QED + Theorem fml_rel_check_cstep_list: fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ @@ -2380,8 +2401,8 @@ Proof strip_tac>>fs[]>> fs[ind_rel_def]>> rw[]>> - (* TODO should be an easy induction *) - `vimap_rel fmlls' vimap` by cheat>> + `vimap_rel fmlls' vimap` by + metis_tac[vimap_rel_core_from_inds]>> metis_tac[IS_SOME_EXISTS,option_CLAUSES]) >- ( (* StrengthenToCore *) gvs[check_cstep_list_def,AllCaseEqs(),check_cstep_def]>> @@ -2434,7 +2455,7 @@ Proof fs[ind_rel_def]>> rw[]>> metis_tac[IS_SOME_EXISTS,option_CLAUSES])>> - (*easy *) cheat) + metis_tac[vimap_rel_core_from_inds]) >- ( (* UnloadOrder *) gvs[check_cstep_list_def,AllCaseEqs(),check_cstep_def]) >- ( (* StoreOrder *) From 80250d614cb9bb4b9a25dad8e101a747b20b0a42 Mon Sep 17 00:00:00 2001 From: Magnus Myreen Date: Sun, 28 Jan 2024 20:34:10 +0100 Subject: [PATCH 18/38] Remove a cheat --- .../pseudo_bool/array/npbc_listScript.sml | 53 ++++++++++++++++++- 1 file changed, 52 insertions(+), 1 deletion(-) diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 3d8d3b5071..a567e69c27 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -1828,12 +1828,63 @@ Proof simp[any_el_ALT] QED +Theorem lookup_update_vimap: + ∀v0 vimap i ls x. + sptree$lookup x vimap = SOME ls ∧ MEM i ls ⇒ + ∃ls. lookup x (update_vimap vimap n v0) = SOME ls ∧ MEM i ls +Proof + Induct \\ gvs [update_vimap_def,FORALL_PROD] \\ rw [] + \\ last_x_assum irule + \\ gvs [lookup_insert] \\ rw [] + \\ Cases_on ‘lookup p_2 vimap’ \\ gvs [opt_cons_def] +QED + +Theorem lookup_update_vimap_MEM: + ∀v0 vimap. + MEM x (MAP SND v0) ⇒ + ∃ls. lookup x (update_vimap vimap i v0) = SOME ls ∧ MEM i ls +Proof + Induct \\ gvs [update_vimap_def,FORALL_PROD] \\ reverse (rw []) + >- (last_x_assum irule \\ fs []) + \\ irule lookup_update_vimap + \\ gvs [lookup_insert] + \\ Cases_on ‘lookup p_2 vimap’ \\ gvs [opt_cons_def] +QED + +Theorem vimap_rel_LUPDATE: + n < LENGTH fml ∧ + vimap_rel fml vimap ⇒ + vimap_rel (LUPDATE (SOME (v,b)) n fml) + (update_vimap vimap n (FST v)) +Proof + gvs [vimap_rel_def,EL_LUPDATE] + \\ rpt strip_tac + \\ PairCases_on ‘v’ \\ gvs [] + \\ Cases_on ‘i = n’ \\ gvs [] + >- (irule lookup_update_vimap_MEM \\ fs []) + \\ first_x_assum drule_all \\ strip_tac \\ gvs [] + \\ irule lookup_update_vimap \\ fs [] +QED + +Theorem vimap_rel_nones: + vimap_rel fml vimap ⇒ + vimap_rel (fml ++ REPLICATE n NONE) vimap +Proof + rw [vimap_rel_def] + \\ last_x_assum irule + \\ Cases_on ‘i < LENGTH fml’ \\ gvs [EL_APPEND1] + \\ gvs [EL_APPEND2,NOT_LESS] + \\ gvs [EL_REPLICATE] +QED + Theorem vimap_rel_update_resize_update_vimap: vimap_rel fml vimap ⇒ vimap_rel (update_resize fml NONE (SOME (v,b)) n) (update_vimap vimap n (FST v)) Proof - cheat + rewrite_tac [update_resize_def] \\ rw [] + \\ irule vimap_rel_LUPDATE \\ fs [] + \\ irule vimap_rel_nones \\ fs [] QED Theorem opt_update_inds_vimap_rel: From 4115cbbda674ea2239c5c4fd050f8ea0701e418a Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Mon, 29 Jan 2024 09:57:34 +0800 Subject: [PATCH 19/38] remove cheat in npbc_list --- .../pseudo_bool/array/npbc_listScript.sml | 77 ++++++++++++++++--- 1 file changed, 67 insertions(+), 10 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index a567e69c27..45de08ff0f 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -1430,6 +1430,66 @@ Proof every_case_tac>>fs[] QED *) +Theorem subst_opt_aux_MEM: + ∀c old new k. + subst_opt_aux f c = (old,new,k,F) ⇒ + ∃x. MEM x (MAP SND c) ∧ IS_SOME (f x) +Proof + Induct>> simp[npbcTheory.subst_opt_aux_def]>> + Cases>> + rw[npbcTheory.subst_opt_aux_def]>> + rpt (pairarg_tac>>fs[])>>gvs[]>> + every_case_tac>>gvs[]>> + metis_tac[IS_SOME_EXISTS] +QED + +Theorem IS_SOME_subst_opt: + IS_SOME (subst_opt f c) ⇒ + ∃x. MEM x (MAP SND (FST c)) ∧ IS_SOME (f x) +Proof + Cases_on`c`>>rw[npbcTheory.subst_opt_eq]>> + rpt (pairarg_tac>>fs[])>> + every_case_tac>>gvs[]>> + metis_tac[subst_opt_aux_MEM] +QED + +Theorem mk_subst_cases: + mk_subst s = + case s of + [(n,v)] => INL (n,v) + | _ => INR (spt_to_vec (fromAList s)) +Proof + every_case_tac>>fs[mk_subst_def] +QED + +Theorem MEM_get_indices_mk_subst: + vimap_rel fmlls vimap ∧ + ind_rel fmlls inds ∧ + any_el i fmlls NONE = SOME (c,b) ∧ + IS_SOME (subst_opt (subst_fun (mk_subst s)) c) + ⇒ + MEM i (get_indices fmlls inds (mk_subst s) vimap) +Proof + rw[get_indices_def]>> + TOP_CASE_TAC>>rw[] + >- ( + drule IS_SOME_subst_opt>>simp[subst_fun_def]>> + gvs[vimap_rel_def,any_el_ALT]>>rw[]>> + first_x_assum drule>> + disch_then drule>> + simp[]>> + disch_then drule>> + rw[]>>gvs[IS_SOME_EXISTS,AllCaseEqs()]>> + gvs[reindex_characterize,MEM_FILTER]>> + gvs[IS_SOME_EXISTS,any_el_ALT]) + >- ( + drule IS_SOME_subst_opt>> + gvs[mk_subst_cases]>>every_case_tac>> + gvs[EVAL``length (spt_to_vec LN)``,subst_fun_def]) + >- ( + gvs[reindex_characterize,MEM_FILTER]>> + gvs[ind_rel_def]) +QED Theorem fml_rel_fml_rel_vimap_rel: fml_rel fml fmlls ∧ @@ -1529,21 +1589,18 @@ Proof rw[]>>eq_tac>>rw[] >- ( fs[MEM_toAList,lookup_map_opt,AllCaseEqs()]>> - (* Prove that every goal is covered, - use the fact that subst_opt ... = SOME - and earliest_rel *) - cheat - (*match_mp_tac (GEN_ALL MEM_subst_indexes)>> + match_mp_tac (GEN_ALL MEM_subst_indexes)>> gvs[lookup_mk_core_fml]>> - first_x_assum (irule_at Any)>> + first_assum (irule_at Any)>> `∃b'. lookup p_1 fml = SOME (x',b') ∧ (b ⇒ b')` by ( gvs[lookup_core_only_def,AllCaseEqs()])>> CONJ_TAC>- ( - drule FST_reindex_characterize>> - simp[MEM_FILTER]>> - fs[fml_rel_def,ind_rel_def])>> + match_mp_tac (GEN_ALL MEM_get_indices_mk_subst)>> + gvs[IS_SOME_EXISTS]>> + first_x_assum (irule_at Any)>> + metis_tac[fml_rel_def])>> simp[rollback_def,lookup_core_only_list_list_delete_list,MEM_MAP,MEM_COUNT_LIST]>> rw[] >- ( @@ -1562,7 +1619,7 @@ Proof drule (GSYM fml_rel_lookup_core_only)>> strip_tac>>fs[]>> gvs[lookup_core_only_list_def,AllCaseEqs()]>> - metis_tac[])*))>> + metis_tac[]))>> drule subst_indexes_MEM>> rw[MEM_toAList,lookup_map_opt]>> gvs[reindex_characterize]>> From f24ab50f1428cc403b12b1dc6f78f466d64b0aff Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Mon, 29 Jan 2024 18:09:53 +0800 Subject: [PATCH 20/38] modify obj subst --- .../array/npbc_arrayProgScript.sml | 107 +++++---- .../pseudo_bool/array/npbc_listScript.sml | 212 +++++++++++++++--- .../array/npbc_parseProgScript.sml | 44 ++-- 3 files changed, 261 insertions(+), 102 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_arrayProgScript.sml b/examples/pseudo_bool/array/npbc_arrayProgScript.sml index 2ad728f1a4..543098e862 100644 --- a/examples/pseudo_bool/array/npbc_arrayProgScript.sml +++ b/examples/pseudo_bool/array/npbc_arrayProgScript.sml @@ -1548,13 +1548,14 @@ QED val res = translate npbc_checkTheory.extract_pids_def; Definition do_rso_def: - do_rso ord s c obj = - red_subgoals ord (subst_fun s) c obj + do_rso ord s c obj vomap = + fast_red_subgoals ord s c obj vomap End val res = translate list_list_insert_def; val res = translate npbcTheory.dom_subst_def; -val res = translate npbc_checkTheory.red_subgoals_def; +val res = translate fast_obj_constraint_def; +val res = translate fast_red_subgoals_def; val res = translate do_rso_def; val res = translate npbc_checkTheory.list_pair_eq_def; @@ -1991,6 +1992,9 @@ val get_indices_arr = process_topdecs` Overload "vimap_TYPE" = `` SPTREE_SPT_TYPE (LIST_TYPE NUM)`` +Overload "vomap_TYPE" = `` + SPTREE_SPT_TYPE UNIT_TYPE`` + Theorem get_indices_arr_spec: LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ (LIST_TYPE NUM) inds indsv ∧ @@ -2042,7 +2046,7 @@ End val res = translate print_lno_mini_def; *) val check_red_arr = process_topdecs` - fun check_red_arr lno ord obj b tcb fml inds id c s pfs idopt vimap = + fun check_red_arr lno ord obj b tcb fml inds id c s pfs idopt vimap vomap = let val s = mk_subst s in case red_fast s idopt pfs of None => @@ -2053,7 +2057,7 @@ val check_red_arr = process_topdecs` case set_indices inds s vimap rinds of (inds',vimap') => let val nc = not_1 c - val rsubs = do_rso ord s c obj + val rsubs = do_rso ord s c obj vomap val cpfs = extract_clauses_arr lno s b fml rsubs pfs [] val fml_not_c = Array.updateResize fml None id (Some (nc,b)) in case check_subproofs_arr lno cpfs b fml_not_c id (id+1) of @@ -2175,12 +2179,13 @@ Theorem check_red_arr_spec: subst_raw_TYPE s sv ∧ pfs_TYPE pfs pfsv ∧ OPTION_TYPE NUM idopt idoptv ∧ - vimap_TYPE vimap vimapv + vimap_TYPE vimap vimapv ∧ + vomap_TYPE vomap vomapv ⇒ app (p : 'ffi ffi_proj) ^(fetch_v "check_red_arr" (get_ml_prog_state())) [lnov; ordv; objv; bv; tcbv; fmlv; indsv; idv; - cv; sv; pfsv; idoptv; vimapv] + cv; sv; pfsv; idoptv; vimapv; vomapv] (ARRAY fmlv fmllsv) (POSTve (λv. @@ -2188,7 +2193,7 @@ Theorem check_red_arr_spec: ARRAY fmlv' fmllsv' * &( case check_red_list ord obj b tcb fmlls inds id - c s pfs idopt vimap of NONE => F + c s pfs idopt vimap vomap of NONE => F | SOME res => PAIR_TYPE (λl v. LIST_REL (OPTION_TYPE bconstraint_TYPE) l fmllsv' ∧ @@ -2201,7 +2206,7 @@ Theorem check_red_arr_spec: ARRAY fmlv' fmllsv' * & (Fail_exn e ∧ check_red_list ord obj b tcb fmlls inds id - c s pfs idopt vimap = NONE))) + c s pfs idopt vimap vomap = NONE))) Proof rw[]>> xcf "check_red_arr" (get_ml_prog_state ())>> @@ -2235,13 +2240,13 @@ Proof (λv. ARRAY aa vv * &(case extract_clauses_list ss b fmlls - (do_rso ord ss c obj) pfs [] of + (do_rso ord ss c obj vomap) pfs [] of NONE => F | SOME res => LIST_TYPE (PAIR_TYPE (OPTION_TYPE (PAIR_TYPE (LIST_TYPE constraint_TYPE) NUM)) (LIST_TYPE NPBC_CHECK_LSTEP_TYPE)) res v)) (λe. ARRAY aa vv * & (Fail_exn e ∧ - extract_clauses_list ss b fmlls (do_rso ord ss c obj) pfs [] = NONE)))` + extract_clauses_list ss b fmlls (do_rso ord ss c obj vomap) pfs [] = NONE)))` >- ( xapp>>xsimpl>> fs[LIST_TYPE_def]>> @@ -2313,7 +2318,8 @@ val res = translate opt_cons_def; val res = translate update_vimap_def; val check_sstep_arr = process_topdecs` - fun check_sstep_arr lno step ord obj tcb fml inds id vimap = + fun check_sstep_arr lno step ord obj tcb fml inds id + vimap vomap = case step of Lstep lstep => (case check_lstep_arr lno lstep False fml 0 id of @@ -2327,7 +2333,7 @@ val check_sstep_arr = process_topdecs` id'+1))) )) | Red c s pfs idopt => (case check_red_arr lno ord obj False tcb - fml inds id c s pfs idopt vimap of + fml inds id c s pfs idopt vimap vomap of (fml',(rinds,(vimap',id'))) => (Array.updateResize fml' None id' (Some (c,tcb)), @@ -2347,11 +2353,12 @@ Theorem check_sstep_arr_spec: LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ (LIST_TYPE NUM) inds indsv ∧ NUM id idv ∧ - vimap_TYPE vimap vimapv + vimap_TYPE vimap vimapv ∧ + vomap_TYPE vomap vomapv ⇒ app (p : 'ffi ffi_proj) ^(fetch_v "check_sstep_arr" (get_ml_prog_state())) - [lnov; stepv; ordv; objv; tcbv; fmlv; indsv; idv; vimapv] + [lnov; stepv; ordv; objv; tcbv; fmlv; indsv; idv; vimapv; vomapv] (ARRAY fmlv fmllsv) (POSTve (λv. @@ -2359,7 +2366,7 @@ Theorem check_sstep_arr_spec: ARRAY fmlv' fmllsv' * &( case check_sstep_list step ord obj tcb - fmlls inds id vimap of NONE => F + fmlls inds id vimap vomap of NONE => F | SOME res => PAIR_TYPE (λl v. LIST_REL (OPTION_TYPE bconstraint_TYPE) l fmllsv' ∧ @@ -2372,7 +2379,7 @@ Theorem check_sstep_arr_spec: ARRAY fmlv' fmllsv' * & (Fail_exn e ∧ check_sstep_list step ord obj tcb - fmlls inds id vimap = NONE))) + fmlls inds id vimap vomap = NONE))) Proof rw[]>> xcf "check_sstep_arr" (get_ml_prog_state ())>> @@ -2431,14 +2438,14 @@ Proof xmatch>> xlet_autop>> rename1`check_red_list ord obj F tcb fmlls inds id - c s pfs idopt vimap`>> + c s pfs idopt vimap vomap`>> xlet`POSTve (λv. SEP_EXISTS fmlv' fmllsv'. ARRAY fmlv' fmllsv' * &( case check_red_list ord obj F tcb fmlls inds id - c s pfs idopt vimap of NONE => F + c s pfs idopt vimap vomap of NONE => F | SOME res => PAIR_TYPE (λl v. LIST_REL (OPTION_TYPE bconstraint_TYPE) l fmllsv' ∧ @@ -2450,7 +2457,7 @@ Proof ARRAY fmlv' fmllsv' * & (Fail_exn e ∧ check_red_list ord obj F tcb fmlls inds id - c s pfs idopt vimap = NONE))` + c s pfs idopt vimap vomap = NONE))` >- ( xapp >> xsimpl>> CONJ_TAC >- EVAL_TAC>> @@ -3228,8 +3235,11 @@ val res = translate err_obj_check_string_def; val res = translate npbc_checkTheory.eq_obj_def; val res = translate npbc_checkTheory.check_eq_obj_def; +val res = translate list_to_num_set_def; +val res = translate mk_vomap_def; + val check_cstep_arr = process_topdecs` - fun check_cstep_arr lno cstep fml inds vimap pc = + fun check_cstep_arr lno cstep fml inds vimap vomap pc = case cstep of Dom c s pfs idopt => ( case get_ord pc of None => @@ -3240,12 +3250,13 @@ val check_cstep_arr = process_topdecs` (fml',(rinds,id')) => (Array.updateResize fml' None id' (Some (c,get_tcb pc)), (sorted_insert id' rinds, - (update_vimap vimap id' (fst c), set_id pc (id'+1))))) + (update_vimap vimap id' (fst c), + (vomap, set_id pc (id'+1)))))) | Sstep sstep => ( case check_sstep_arr lno sstep (get_ord pc) (get_obj pc) - (get_tcb pc) fml inds (get_id pc) vimap of + (get_tcb pc) fml inds (get_id pc) vimap vomap of (fml',(inds',(vimap',id'))) => - (fml',(inds',(vimap',set_id pc id')))) + (fml',(inds',(vimap',(vomap, set_id pc id'))))) | Checkeddelete n s pfs idopt => ( if check_tcb_idopt_pc pc idopt then @@ -3256,35 +3267,36 @@ val check_cstep_arr = process_topdecs` case check_red_arr lno (get_ord pc) (get_obj pc) True (get_tcb pc) fml inds (get_id pc) - c s pfs idopt vimap of (fml',(inds',(vimap',id'))) => - (fml',(inds',(vimap',set_id pc id'))))) + c s pfs idopt vimap vomap of + (fml',(inds',(vimap',id'))) => + (fml',(inds',(vimap',(vomap,set_id pc id')))))) else raise Fail (format_failure lno "invalid proof state for checked deletion")) | Uncheckeddelete ls => ( if check_tcb_ord pc then (list_delete_arr ls fml; - (fml, (inds, (vimap, set_chk pc False)))) + (fml, (inds, (vimap, (vomap, set_chk pc False))))) else case all_core_arr fml inds [] of None => raise Fail (format_failure lno "not all constraints in core") | Some inds' => (list_delete_arr ls fml; - (fml, (inds', (vimap, set_chk pc False))))) + (fml, (inds', (vimap, (vomap, set_chk pc False)))))) | Transfer ls => ( let val fml' = core_from_inds_arr lno fml ls in - (fml', (inds, (vimap, pc))) + (fml', (inds, (vimap, (vomap, pc)))) end) | Strengthentocore b => ( let val inds' = reindex_arr fml inds in if b then ( let val fml' = core_from_inds_arr lno fml inds' in - (fml', (inds', (vimap, set_tcb pc b))) + (fml', (inds', (vimap, (vomap, set_tcb pc b)))) end) else - (fml,(inds', (vimap, set_tcb pc b))) + (fml,(inds', (vimap, (vomap, set_tcb pc b)))) end) | Loadorder nn xs => let val inds' = reindex_arr fml inds in @@ -3294,7 +3306,7 @@ val check_cstep_arr = process_topdecs` | Some ord' => if List.length xs = List.length (fst (snd ord')) then let val fml' = core_from_inds_arr lno fml inds' in - (fml', (inds', (vimap, set_ord pc (Some (ord',xs))))) + (fml', (inds', (vimap, (vomap, set_ord pc (Some (ord',xs)))))) end else raise Fail @@ -3304,11 +3316,11 @@ val check_cstep_arr = process_topdecs` (case get_ord pc of None => raise Fail (format_failure lno ("no order loaded")) | Some spo => - (fml,(inds,(vimap,set_ord pc None)))) + (fml,(inds,(vimap,(vomap,set_ord pc None))))) | Storeorder name spo ws pfsr pfst => if check_storeorder spo ws pfst pfsr then - (fml, (inds, (vimap, set_orders pc ((name,spo)::get_orders pc)))) + (fml, (inds, (vimap, (vomap, set_orders pc ((name,spo)::get_orders pc))))) else raise Fail (format_failure lno ("invalid order definition")) | Obj w mi bopt => ( @@ -3329,20 +3341,21 @@ val check_cstep_arr = process_topdecs` (Array.updateResize fml None id (Some (c,True)), (sorted_insert id inds, (update_vimap vimap id (fst c), - obj_update pc (id+1) bound' dbound'))) + (vomap, + obj_update pc (id+1) bound' dbound')))) end else - (fml, (inds, (vimap, obj_update pc (get_id pc) bound' dbound'))) + (fml, (inds, (vimap, (vomap, obj_update pc (get_id pc) bound' dbound')))) end ) | Changeobj b fc' pfs => (case check_change_obj_arr lno b fml (get_id pc) (get_obj pc) fc' pfs of (fml',(fc',id')) => - (fml', (inds, (vimap, change_obj_update pc id' fc')))) + (fml', (inds, (vimap, (mk_vomap fc', change_obj_update pc id' fc'))))) | Checkobj fc' => if check_eq_obj (get_obj pc) fc' - then (fml, (inds, (vimap, pc))) + then (fml, (inds, (vimap, (vomap, pc)))) else raise Fail (format_failure lno (err_obj_check_string (get_obj pc) fc')) @@ -3376,18 +3389,19 @@ Theorem check_cstep_arr_spec: LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ (LIST_TYPE NUM) inds indsv ∧ NPBC_CHECK_PROOF_CONF_TYPE pc pcv ∧ - vimap_TYPE vimap vimapv + vimap_TYPE vimap vimapv ∧ + vomap_TYPE vomap vomapv ⇒ app (p : 'ffi ffi_proj) ^(fetch_v "check_cstep_arr" (get_ml_prog_state())) - [lnov; cstepv; fmlv; indsv; vimapv; pcv] + [lnov; cstepv; fmlv; indsv; vimapv; vomapv; pcv] (ARRAY fmlv fmllsv) (POSTve (λv. SEP_EXISTS fmlv' fmllsv'. ARRAY fmlv' fmllsv' * &( - case check_cstep_list cstep fmlls inds vimap pc of + case check_cstep_list cstep fmlls inds vimap vomap pc of NONE => F | SOME res => PAIR_TYPE (λl v. @@ -3395,14 +3409,15 @@ Theorem check_cstep_arr_spec: v = fmlv') (PAIR_TYPE (LIST_TYPE NUM) (PAIR_TYPE (vimap_TYPE) - NPBC_CHECK_PROOF_CONF_TYPE)) + (PAIR_TYPE (vomap_TYPE) + NPBC_CHECK_PROOF_CONF_TYPE))) res v )) (λe. SEP_EXISTS fmlv' fmllsv'. ARRAY fmlv' fmllsv' * & (Fail_exn e ∧ - check_cstep_list cstep fmlls inds vimap pc = NONE))) + check_cstep_list cstep fmlls inds vimap vomap pc = NONE))) Proof rw[]>> xcf "check_cstep_arr" (get_ml_prog_state ())>> @@ -3487,14 +3502,14 @@ Proof rpt xlet_autop>> fs[get_id_def,get_tcb_def,get_obj_def,get_ord_def]>> rename1`check_red_list ord obj T pc.tcb (delete_list n fmlls) - inds pc.id c s pfs idopt vimap`>> + inds pc.id c s pfs idopt vimap vomap`>> xlet`POSTve (λv. SEP_EXISTS fmlv' fmllsv'. ARRAY fmlv' fmllsv' * &( case check_red_list ord obj T pc.tcb (delete_list n fmlls) - inds pc.id c s pfs idopt vimap of NONE => F + inds pc.id c s pfs idopt vimap vomap of NONE => F | SOME res => PAIR_TYPE (λl v. LIST_REL (OPTION_TYPE bconstraint_TYPE) l fmllsv' ∧ @@ -3506,7 +3521,7 @@ Proof ARRAY fmlv' fmllsv' * & (Fail_exn e ∧ check_red_list ord obj T pc.tcb (delete_list n fmlls) - inds pc.id c s pfs idopt vimap = NONE))` + inds pc.id c s pfs idopt vimap vomap = NONE))` >- ( xapp>>xsimpl>> CONJ_TAC >- EVAL_TAC>> diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 45de08ff0f..583ec68f78 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -774,8 +774,32 @@ Definition set_indices_def: (inds, sptree$insert n rinds vimap) End +(* Fast substitution for obj_constraint if it is *) +Definition fast_obj_constraint_def: + fast_obj_constraint s l vomap = + case s of + INR v => + if length v = 0 then ([],0) + else obj_constraint (subst_fun s) l + | INL (n,_) => + case sptree$lookup n vomap of + NONE => ([],0) + | SOME () => obj_constraint (subst_fun s) l +End + +Definition fast_red_subgoals_def: + fast_red_subgoals ord s def obj vomap = + let cobj = + case obj of NONE => [] + | SOME l => [[not (fast_obj_constraint s l vomap)]] in + let s = subst_fun s in + let c0 = subst s def in (**) + [not c0]::(MAP (λc. [not c]) (dom_subst s ord)) ++ cobj +End + Definition check_red_list_def: - check_red_list ord obj b tcb fml inds id c s pfs idopt vimap = + check_red_list ord obj b tcb fml inds id c s pfs idopt + vimap vomap = let s = mk_subst s in case red_fast s idopt pfs of NONE => ( @@ -783,7 +807,7 @@ Definition check_red_list_def: let (inds',vimap') = set_indices inds s vimap rinds in let nc = not c in let fml_not_c = update_resize fml NONE (SOME (nc,b)) id in - let rsubs = red_subgoals ord (subst_fun s) c obj in + let rsubs = fast_red_subgoals ord s c obj vomap in case extract_clauses_list s b fml rsubs pfs [] of NONE => NONE | SOME cpfs => @@ -846,7 +870,7 @@ End Definition check_sstep_list_def: (check_sstep_list (sstep:sstep) ord obj tcb (fml: (npbc # bool) option list) (inds:num list) (id:num) - vimap = + vimap vomap = case sstep of | Lstep lstep => (case check_lstep_list lstep F fml 0 id of NONE => NONE @@ -854,7 +878,7 @@ Definition check_sstep_list_def: SOME (opt_update_inds rfml c id' inds vimap)) | Red c s pfs idopt => case check_red_list ord obj F tcb fml inds id c s pfs - idopt vimap of + idopt vimap vomap of SOME (rfml,rinds,vimap',id') => SOME ( update_resize rfml NONE (SOME (c,tcb)) id', @@ -1532,13 +1556,84 @@ Proof simp[any_el_ALT] QED +Definition vomap_rel_def: + vomap_rel obj vomap ⇔ + case obj of + NONE => T + | SOME l => + set (MAP SND (FST l)) = domain vomap +End + +Theorem add_lists_map_negate_coeff: + ∀ls rs. + rs = (MAP (λ(c,l). (-c,l)) ls) ⇒ + add_lists ls rs = ([],SUM (MAP (λi. Num (ABS (FST i))) ls)) +Proof + ho_match_mp_tac npbcTheory.add_lists_ind>> + simp[npbcTheory.add_lists_def,npbcTheory.add_terms_def] +QED + +Theorem subst_aux_id: + ∀l. + EVERY (\v. f v = NONE) (MAP SND l) ⇒ + subst_aux f l = (l,[],0) +Proof + Induct>-simp[npbcTheory.subst_aux_def]>> + Cases>> + rw[npbcTheory.subst_aux_def] +QED + +Theorem subst_lhs_id: + EVERY (\v. f v = NONE) (MAP SND l) ⇒ + subst_lhs f l = (l, 0) +Proof + rw[npbcTheory.subst_lhs_def]>> + rpt(pairarg_tac>>fs[])>> + drule subst_aux_id>>strip_tac>> + gvs[EVAL``clean_up []``]>> + Cases_on`l`>> + gvs[npbcTheory.add_lists_def] +QED + +Theorem vomap_rel_fast_obj_constraint: + vomap_rel (SOME l) vomap ⇒ + fast_obj_constraint s l vomap = + obj_constraint (subst_fun s) l +Proof + rw[fast_obj_constraint_def]>> + every_case_tac>> + Cases_on`l`>> + fs[npbcTheory.obj_constraint_def,subst_fun_def]>> + rpt (pairarg_tac>>fs[])>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[add_lists_map_negate_coeff]>>rw[]>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[subst_lhs_id]>> + fs[vomap_rel_def]>> + simp[EVERY_MAP,LAMBDA_PROD,subst_fun_def]>> + gvs[EVERY_MEM]>> + rw[]>>pairarg_tac>>fs[EXTENSION,MEM_MAP,domain_lookup]>> + metis_tac[option_CLAUSES,SND,PAIR] +QED + +Theorem vomap_rel_fast_red_subgoals: + vomap_rel obj vomap ⇒ + fast_red_subgoals ord s def obj vomap = + red_subgoals ord (subst_fun s) def obj +Proof + rw[fast_red_subgoals_def,red_subgoals_def]>> + every_case_tac>>fs[]>> + metis_tac[vomap_rel_fast_obj_constraint] +QED + Theorem fml_rel_check_red_list: fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ vimap_rel fmlls vimap ∧ + vomap_rel obj vomap ∧ (∀n. n ≥ id ⇒ any_el n fmlls NONE = NONE) ∧ check_red_list ord obj b tcb fmlls inds id c s pfs - idopt vimap = + idopt vimap vomap = SOME (fmlls', inds', vimap', id') ⇒ check_red ord obj b tcb fml id c s pfs idopt = SOME id' ∧ fml_rel fml fmlls' ∧ @@ -1551,6 +1646,7 @@ Proof fs[check_red_list_def]>> gvs[AllCaseEqs()] >- ( + gvs[vomap_rel_fast_red_subgoals]>> pairarg_tac>>fs[]>> every_case_tac>>gvs[]>> simp[check_red_def]>> @@ -1959,8 +2055,9 @@ Theorem fml_rel_check_sstep_list: fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ vimap_rel fmlls vimap ∧ + vomap_rel obj vomap ∧ (∀n. n ≥ id ⇒ any_el n fmlls NONE = NONE) ∧ - check_sstep_list sstep ord obj tcb fmlls inds id vimap = + check_sstep_list sstep ord obj tcb fmlls inds id vimap vomap = SOME (fmlls',inds',vimap',id') ⇒ ∃fml'. check_sstep sstep ord obj tcb fml id = SOME(fml',id') ∧ @@ -2125,8 +2222,20 @@ Definition check_change_obj_list_def: else NONE) End +Definition mk_vomap_def: + mk_vomap (f,c) = + list_to_num_set (MAP SND f) +End + +Theorem vomap_rel_mk_vomap: + vomap_rel (SOME fc) (mk_vomap fc) +Proof + Cases_on`fc`>>rw[vomap_rel_def,mk_vomap_def]>> + fs[EXTENSION,domain_list_to_num_set] +QED + Definition check_cstep_list_def: - check_cstep_list cstep fml inds vimap pc = + check_cstep_list cstep fml inds vimap vomap pc = case cstep of Dom c s pfs idopt => (case pc.ord of @@ -2152,13 +2261,14 @@ Definition check_cstep_list_def: update_resize rfml NONE (SOME (c,pc.tcb)) id', sorted_insert id' rinds, update_vimap vimap id' (FST c), + vomap, pc with id := id'+1) else NONE))) | Sstep sstep => (case check_sstep_list sstep pc.ord pc.obj pc.tcb - fml inds pc.id vimap of + fml inds pc.id vimap vomap of SOME(fml',inds',vimap',id') => - SOME(fml',inds', vimap',pc with id := id') + SOME(fml',inds', vimap', vomap, pc with id := id') | NONE => NONE) | CheckedDelete n s pfs idopt => ( if check_tcb_idopt pc.tcb idopt then @@ -2167,9 +2277,9 @@ Definition check_cstep_list_def: | SOME c => (let nfml = delete_list n fml in case check_red_list pc.ord pc.obj T pc.tcb - nfml inds pc.id c s pfs idopt vimap of + nfml inds pc.id c s pfs idopt vimap vomap of SOME (ncf',inds',vimap',id') => - SOME (ncf', inds', vimap', pc with <| id := id' |>) + SOME (ncf', inds', vimap', vomap, pc with <| id := id' |>) | NONE => NONE) ) else NONE) | UncheckedDelete ls => ( @@ -2177,25 +2287,25 @@ Definition check_cstep_list_def: if ¬pc.tcb ∧ pc.ord = NONE then SOME (list_delete_list ls fml, inds, - vimap, pc with chk := F) + vimap, vomap, pc with chk := F) else case all_core_list fml inds [] of NONE => NONE | SOME inds' => SOME (list_delete_list ls fml, inds', - vimap, pc with chk := F)) + vimap, vomap, pc with chk := F)) | Transfer ls => (case core_from_inds fml ls of NONE => NONE | SOME fml' => - SOME (fml', inds, vimap, pc)) + SOME (fml', inds, vimap, vomap, pc)) | StrengthenToCore b => (let inds' = reindex fml inds in let pc' = pc with tcb := b in if b then (case core_from_inds fml inds' of NONE => NONE - | SOME fml' => SOME (fml',inds', vimap, pc')) + | SOME fml' => SOME (fml',inds', vimap, vomap, pc')) else - SOME (fml,inds',vimap,pc')) + SOME (fml,inds',vimap, vomap, pc')) | LoadOrder nn xs => (let inds' = reindex fml inds in case ALOOKUP pc.orders nn of NONE => NONE @@ -2203,12 +2313,13 @@ Definition check_cstep_list_def: if LENGTH xs = LENGTH (FST (SND ord')) then case core_from_inds fml inds' of NONE => NONE | SOME fml' => - SOME (fml',inds',vimap,pc with ord := SOME (ord',xs)) + SOME (fml',inds', + vimap,vomap,pc with ord := SOME (ord',xs)) else NONE) | UnloadOrder => (case pc.ord of NONE => NONE | SOME spo => - SOME (fml,inds, vimap, pc with ord := NONE)) + SOME (fml,inds, vimap, vomap, pc with ord := NONE)) | StoreOrder nn spo ws pfsr pfst => if check_good_ord spo ∧ check_ws spo ws then @@ -2216,7 +2327,7 @@ Definition check_cstep_list_def: | SOME id => if check_reflexivity spo pfsr id then SOME (fml, inds, - vimap, pc with orders := (nn,spo)::pc.orders) + vimap, vomap, pc with orders := (nn,spo)::pc.orders) else NONE else NONE @@ -2233,12 +2344,13 @@ Definition check_cstep_list_def: update_resize fml NONE (SOME (c,T)) pc.id, sorted_insert pc.id inds, update_vimap vimap pc.id (FST c), + vomap, pc with <| id := pc.id+1; bound := bound'; dbound := dbound' |>) else - SOME (fml, inds, vimap, + SOME (fml, inds, vimap, vomap, pc with <| bound := bound'; dbound := dbound' |>)) @@ -2247,11 +2359,11 @@ Definition check_cstep_list_def: NONE => NONE | SOME (fml',fc',id') => SOME ( - fml', inds, vimap, + fml', inds, vimap, mk_vomap fc', pc with <| id:=id'; obj:=SOME fc' |>)) | CheckObj fc' => if check_eq_obj pc.obj fc' - then SOME (fml, inds, vimap, pc) + then SOME (fml, inds, vimap, vomap, pc) else NONE End @@ -2398,14 +2510,16 @@ Theorem fml_rel_check_cstep_list: fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ vimap_rel fmlls vimap ∧ + vomap_rel pc.obj vomap ∧ (∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE) ∧ - check_cstep_list cstep fmlls inds vimap pc = - SOME (fmlls',inds',vimap',pc') ⇒ + check_cstep_list cstep fmlls inds vimap vomap pc = + SOME (fmlls',inds',vimap',vomap',pc') ⇒ ∃fml'. check_cstep cstep fml pc = SOME (fml', pc') ∧ fml_rel fml' fmlls' ∧ ind_rel fmlls' inds' ∧ vimap_rel fmlls' vimap' ∧ + vomap_rel pc'.obj vomap' ∧ (∀n. n ≥ pc'.id ⇒ any_el n fmlls' NONE = NONE) ∧ pc.id ≤ pc'.id Proof @@ -2611,6 +2725,8 @@ Proof metis_tac[ind_rel_reindex])>> CONJ_TAC >- metis_tac[fml_rel_fml_rel_vimap_rel]>> + CONJ_TAC >- + metis_tac[vomap_rel_mk_vomap]>> simp[any_el_rollback]) >- ( (* CheckObj *) fs[check_cstep_def,check_cstep_list_def] @@ -2618,28 +2734,31 @@ Proof QED Definition check_csteps_list_def: - (check_csteps_list [] fml inds vimap pc = - SOME (fml, inds, vimap, pc)) ∧ - (check_csteps_list (c::cs) fml inds vimap pc = - case check_cstep_list c fml inds vimap pc of + (check_csteps_list [] fml inds vimap vomap pc = + SOME (fml, inds, vimap, vomap, pc)) ∧ + (check_csteps_list (c::cs) fml inds vimap vomap pc = + case check_cstep_list c fml inds vimap vomap pc of NONE => NONE - | SOME(fml', inds', vimap', pc') => - check_csteps_list cs fml' inds' vimap' pc') + | SOME(fml', inds', vimap', vomap', pc') => + check_csteps_list cs fml' inds' vimap' vomap' pc') End Theorem fml_rel_check_csteps_list: - ∀csteps fml fmlls inds vimap pc fmlls' inds' vimap' pc'. + ∀csteps fml fmlls inds vimap vomap pc + fmlls' inds' vimap' vomap' pc'. fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ vimap_rel fmlls vimap ∧ + vomap_rel pc.obj vomap ∧ (∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE) ∧ - check_csteps_list csteps fmlls inds vimap pc = - SOME (fmlls', inds', vimap', pc') ⇒ + check_csteps_list csteps fmlls inds vimap vomap pc = + SOME (fmlls', inds', vimap', vomap', pc') ⇒ ∃fml'. check_csteps csteps fml pc = SOME (fml', pc') ∧ fml_rel fml' fmlls' ∧ ind_rel fmlls' inds' ∧ vimap_rel fmlls' vimap' ∧ + vomap_rel pc'.obj vomap' ∧ (∀n. n ≥ pc'.id ⇒ any_el n fmlls' NONE = NONE) ∧ pc.id ≤ pc'.id Proof @@ -2870,21 +2989,27 @@ Proof fs[MAP_FST_enumerate,MEM_GENLIST] QED +Definition mk_vomap_opt_def: + (mk_vomap_opt NONE = LN) ∧ + (mk_vomap_opt (SOME fc) = mk_vomap fc) +End + Theorem check_csteps_list_concl: check_csteps_list cs (FOLDL (λacc (i,v). update_resize acc NONE (SOME (v,T)) i) (REPLICATE m NONE) (enumerate 1 fml)) (REVERSE (MAP FST (enumerate 1 fml))) (FOLDL (λacc (i,v). update_vimap acc i (FST v)) LN (enumerate 1 fml)) + (mk_vomap_opt obj) (init_conf (LENGTH fml + 1) chk obj) = - SOME(fmlls',inds',vimap',pc') ∧ + SOME(fmlls',inds',vimap',vomap',pc') ∧ check_hconcl_list fml obj fmlls' pc'.obj pc'.bound pc'.dbound hconcl ⇒ sem_concl (set fml) obj (hconcl_concl hconcl) Proof rw[]>> qmatch_asmsub_abbrev_tac`check_csteps_list cs fmlls inds - vimap pc = _`>> + vimap vomap pc = _`>> `fml_rel (build_fml T 1 fml) fmlls` by simp[Abbr`fmlls`,fml_rel_FOLDL_update_resize]>> `ind_rel fmlls inds` by ( @@ -2895,6 +3020,12 @@ Proof `vimap_rel fmlls vimap` by ( unabbrev_all_tac>> simp[vimap_rel_FOLDL_update_resize])>> + `vomap_rel pc.obj vomap` by ( + unabbrev_all_tac>> + simp[init_conf_def]>> + Cases_on`obj`>- + simp[vomap_rel_def]>> + metis_tac[vomap_rel_mk_vomap,mk_vomap_opt_def])>> `∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE` by ( rw[Abbr`pc`,Abbr`fmlls`,any_el_ALT,init_conf_def]>> DEP_REWRITE_TAC [FOLDL_update_resize_lookup]>> @@ -2991,14 +3122,15 @@ Theorem check_csteps_list_output: (REPLICATE m NONE) (enumerate 1 fml)) (REVERSE (MAP FST (enumerate 1 fml))) (FOLDL (λacc (i,v). update_vimap acc i (FST v)) LN (enumerate 1 fml)) + (mk_vomap_opt obj) (init_conf (LENGTH fml + 1) chk obj) = - SOME(fmlls',inds',vimap',pc') ∧ + SOME(fmlls',inds',vimap',vomap',pc') ∧ check_output_list fmlls' inds' pc'.obj pc'.bound pc'.dbound pc'.chk fmlt objt output ⇒ sem_output (set fml) obj pc'.bound (set fmlt) objt output Proof rw[]>> - qmatch_asmsub_abbrev_tac`check_csteps_list cs fmlls inds vimap pc = _`>> + qmatch_asmsub_abbrev_tac`check_csteps_list cs fmlls inds vimap vomap pc = _`>> `fml_rel (build_fml T 1 fml) fmlls` by simp[Abbr`fmlls`,fml_rel_FOLDL_update_resize]>> `ind_rel fmlls inds` by ( @@ -3009,6 +3141,12 @@ Proof `vimap_rel fmlls vimap` by ( unabbrev_all_tac>> simp[vimap_rel_FOLDL_update_resize])>> + `vomap_rel pc.obj vomap` by ( + unabbrev_all_tac>> + simp[init_conf_def]>> + Cases_on`obj`>- + simp[vomap_rel_def]>> + metis_tac[vomap_rel_mk_vomap,mk_vomap_opt_def])>> `∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE` by ( rw[Abbr`pc`,Abbr`fmlls`,any_el_ALT,init_conf_def]>> DEP_REWRITE_TAC [FOLDL_update_resize_lookup]>> diff --git a/examples/pseudo_bool/array/npbc_parseProgScript.sml b/examples/pseudo_bool/array/npbc_parseProgScript.sml index 730fc577aa..e604eff267 100644 --- a/examples/pseudo_bool/array/npbc_parseProgScript.sml +++ b/examples/pseudo_bool/array/npbc_parseProgScript.sml @@ -2312,15 +2312,15 @@ QED (* returns the necessary information to check the output and conclusion sections *) val check_unsat'' = process_topdecs ` - fun check_unsat'' fns fd lno fml inds vimap pc = + fun check_unsat'' fns fd lno fml inds vimap vomap pc = case parse_cstep fns fd lno of (Inl s, (fns', lno')) => (lno', (s, (fns', (fml, (inds, pc))))) | (Inr cstep, (fns', lno')) => - (case check_cstep_arr lno cstep fml inds vimap pc of - (fml', (inds', (vimap', pc'))) => - check_unsat'' fns' fd lno' fml' inds' vimap' pc')` |> append_prog + (case check_cstep_arr lno cstep fml inds vimap vomap pc of + (fml', (inds', (vimap', (vomap', pc')))) => + check_unsat'' fns' fd lno' fml' inds' vimap' vomap' pc')` |> append_prog Theorem parse_sstep_LENGTH: ∀f ss res f' ss'. @@ -2365,15 +2365,15 @@ QED returning the last encountered state *) Definition parse_and_run_def: parse_and_run fns ss - fml inds vimap pc = + fml inds vimap vomap pc = case parse_cstep fns ss of NONE => NONE | SOME (INL s, fns', rest) => SOME (rest, s, fns', fml, inds, pc) | SOME (INR cstep, fns', rest) => - (case check_cstep_list cstep fml inds vimap pc of - SOME (fml', inds', vimap', pc') => - parse_and_run fns' rest fml' inds' vimap' pc' + (case check_cstep_list cstep fml inds vimap vomap pc of + SOME (fml', inds', vimap', vomap', pc') => + parse_and_run fns' rest fml' inds' vimap' vomap' pc' | res => NONE) Termination WF_REL_TAC `measure (LENGTH o FST o SND)`>> @@ -2412,19 +2412,20 @@ Proof QED Theorem check_unsat''_spec: - ∀fns ss fmlls inds vimap pc - fnsv lno lnov fmllsv indsv pcv lines fs fmlv vimapv. + ∀fns ss fmlls inds vimap vomap pc + fnsv lno lnov fmllsv indsv pcv lines fs fmlv vimapv vomapv. fns_TYPE a fns fnsv ∧ NUM lno lnov ∧ LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ (LIST_TYPE NUM) inds indsv ∧ NPBC_CHECK_PROOF_CONF_TYPE pc pcv ∧ vimap_TYPE vimap vimapv ∧ + vomap_TYPE vomap vomapv ∧ MAP toks_fast lines = ss ⇒ app (p : 'ffi ffi_proj) ^(fetch_v "check_unsat''" (get_ml_prog_state())) - [fnsv; fdv; lnov; fmlv; indsv; vimapv; pcv] + [fnsv; fdv; lnov; fmlv; indsv; vimapv; vomapv; pcv] (STDIO fs * INSTREAM_LINES fd fdv lines fs * ARRAY fmlv fmllsv) (POSTve (λv. @@ -2433,7 +2434,7 @@ Theorem check_unsat''_spec: INSTREAM_LINES fd fdv lines' (forwardFD fs fd k) * ARRAY fmlv' fmllsv' * &( - parse_and_run fns ss fmlls inds vimap pc = + parse_and_run fns ss fmlls inds vimap vomap pc = SOME (MAP toks_fast lines',res) ∧ PAIR_TYPE NUM ( PAIR_TYPE (LIST_TYPE (SUM_TYPE STRING_TYPE INT)) ( @@ -2448,7 +2449,7 @@ Theorem check_unsat''_spec: ARRAY fmlv' fmllsv' * STDIO (forwardFD fs fd k) * INSTREAM_LINES fd fdv lines' (forwardFD fs fd k) * &(Fail_exn e ∧ - parse_and_run fns ss fmlls inds vimap pc = NONE))) + parse_and_run fns ss fmlls inds vimap vomap pc = NONE))) Proof ho_match_mp_tac (fetch "-" "parse_and_run_ind")>> rw[]>> @@ -2927,6 +2928,8 @@ val res = translate init_conf_def; val res = translate hconcl_concl_def; val res = translate conv_boutput_hconcl_def; +val res = translate npbc_listTheory.mk_vomap_opt_def; + val check_unsat' = process_topdecs ` fun check_unsat' fns fd lno fml obj fmlt objt = let @@ -2935,9 +2938,10 @@ val check_unsat' = process_topdecs ` val arr = fill_arr arr 1 fml val inds = rev_enum_full 1 fml val vimap = fold_update_vimap_enum_full 1 fml + val vomap = mk_vomap_opt obj val pc = init_conf id True obj in - (case check_unsat'' fns fd lno arr inds vimap pc of + (case check_unsat'' fns fd lno arr inds vimap vomap pc of (lno', (s, (fns',( (fml', (inds', pc')))))) => conv_boutput_hconcl @@ -2950,10 +2954,10 @@ val check_unsat' = process_topdecs ` end` |> append_prog; Theorem parse_and_run_check_csteps_list: - ∀fns ss fml inds vimap pc rest s fns' fml' inds' pc'. - parse_and_run fns ss fml inds vimap pc = SOME (rest, s, fns', (fml', inds', pc')) ⇒ - ∃csteps vimap'. - check_csteps_list csteps fml inds vimap pc = SOME (fml', inds', vimap', pc') + ∀fns ss fml inds vimap vomap pc rest s fns' fml' inds' pc'. + parse_and_run fns ss fml inds vimap vomap pc = SOME (rest, s, fns', (fml', inds', pc')) ⇒ + ∃csteps vimap' vomap'. + check_csteps_list csteps fml inds vimap vomap pc = SOME (fml', inds', vimap', vomap', pc') Proof ho_match_mp_tac parse_and_run_ind>> rw[]>> @@ -3047,8 +3051,10 @@ Proof xlet_autop>> qmatch_asmsub_abbrev_tac`vimap_TYPE vimap vimapv`>> + qmatch_asmsub_abbrev_tac`vomap_TYPE vomap vomapv`>> Cases_on` parse_and_run fns (MAP toks_fast lines) fmlls inds vimap + vomap (init_conf (LENGTH fml + 1) T obj)` >- ( (* fail to parse and run *) @@ -3103,7 +3109,7 @@ Proof ARRAY fmlv' fmllsv' * &( parse_and_run fns (MAP toks_fast lines) - fmlls inds vimap (init_conf (LENGTH fml + 1) T obj) = + fmlls inds vimap vomap (init_conf (LENGTH fml + 1) T obj) = SOME (MAP toks_fast lines',res) ∧ PAIR_TYPE NUM ( PAIR_TYPE (LIST_TYPE (SUM_TYPE STRING_TYPE INT)) ( From fda6d5934d6c61fc3975f8ff3806a0b6656779fe Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Tue, 30 Jan 2024 00:24:01 +0800 Subject: [PATCH 21/38] update obj mapping --- .../array/npbc_arrayProgScript.sml | 116 +++++++++++++++++- .../pseudo_bool/array/npbc_listScript.sml | 56 ++++++--- .../array/npbc_parseProgScript.sml | 27 +++- 3 files changed, 177 insertions(+), 22 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_arrayProgScript.sml b/examples/pseudo_bool/array/npbc_arrayProgScript.sml index 543098e862..a4d551b2bc 100644 --- a/examples/pseudo_bool/array/npbc_arrayProgScript.sml +++ b/examples/pseudo_bool/array/npbc_arrayProgScript.sml @@ -1992,8 +1992,7 @@ val get_indices_arr = process_topdecs` Overload "vimap_TYPE" = `` SPTREE_SPT_TYPE (LIST_TYPE NUM)`` -Overload "vomap_TYPE" = `` - SPTREE_SPT_TYPE UNIT_TYPE`` +Overload "vomap_TYPE" = ``STRING_TYPE`` Theorem get_indices_arr_spec: LIST_REL (OPTION_TYPE bconstraint_TYPE) fmlls fmllsv ∧ @@ -3235,8 +3234,115 @@ val res = translate err_obj_check_string_def; val res = translate npbc_checkTheory.eq_obj_def; val res = translate npbc_checkTheory.check_eq_obj_def; -val res = translate list_to_num_set_def; -val res = translate mk_vomap_def; +Definition w8z_def: + w8z = (0w: word8) +End + +Definition w8o_def: + w8o = (1w: word8) +End + +val w8z_v_thm = translate w8z_def; +val w8o_v_thm = translate w8o_def; + +val fold_update_resize_bitset = process_topdecs` + fun fold_update_resize_bitset ls acc = + case ls of + [] => acc + | (x::xs) => + if x < Word8Array.length acc + then + (Word8Array.update acc x w8o; + fold_update_resize_bitset xs acc) + else + let + val arr = Word8Array.array (2*x+1) w8z + val u = Word8Array.copy acc 0 (Word8Array.length acc) arr 0 in + (Word8Array.update arr x w8o; + fold_update_resize_bitset xs arr) + end + ` |> append_prog; + +Theorem fold_update_resize_bitset_spec: + ∀ls lsv accv accls. + LIST_TYPE NUM ls lsv + ⇒ + app (p : 'ffi ffi_proj) + ^(fetch_v "fold_update_resize_bitset" (get_ml_prog_state())) + [lsv; accv] + (W8ARRAY accv accls) + (POSTv v. + W8ARRAY v (FOLDL (λacc i. update_resize acc w8z w8o i) accls ls)) +Proof + Induct>> + xcf "fold_update_resize_bitset" (get_ml_prog_state ())>> + gvs[LIST_TYPE_def]>>xmatch + >- ( + xvar>>xsimpl)>> + assume_tac w8o_v_thm>> + assume_tac w8z_v_thm>> + rpt xlet_autop>> + xif + >- ( + xlet_autop>> + xapp>>xsimpl>> + simp[update_resize_def])>> + rpt xlet_autop>> + xapp>>xsimpl>> + simp[update_resize_def] +QED + +val mk_vomap_arr = process_topdecs` + fun mk_vomap_arr n fc = + let + val acc = Word8Array.array n w8z + val f = map_snd (fst fc) + val acc = fold_update_resize_bitset f acc in + Word8Array.substring acc 0 (Word8Array.length acc) + end` |> append_prog; + +Theorem map_foldl_rel: + ∀ls accA accB. + MAP (CHR o w2n) accA = accB ⇒ + MAP (CHR ∘ w2n) + (FOLDL (λacc i. update_resize acc w8z w8o i) accA ls) = + FOLDL (λacc i. update_resize acc #"\^@" #"\^A" i) accB ls +Proof + Induct>>rw[]>> + first_x_assum match_mp_tac>> + rw[update_resize_def,LUPDATE_MAP]>> + EVAL_TAC +QED + +Theorem mk_vomap_arr_spec: + NUM n nv ∧ + (PAIR_TYPE (LIST_TYPE (PAIR_TYPE INT NUM)) INT) fc fcv + ⇒ + app (p : 'ffi ffi_proj) + ^(fetch_v "mk_vomap_arr" (get_ml_prog_state())) + [nv; fcv] + (emp) + (POSTv v. &(vomap_TYPE (mk_vomap n fc) v)) +Proof + rw[]>> + xcf "mk_vomap_arr" (get_ml_prog_state ())>> + assume_tac w8z_v_thm>> + xlet_auto>> + rpt xlet_autop>> + xlet_auto>> + rpt xlet_autop>> + xapp>>xsimpl>> + first_x_assum (irule_at Any)>> rw[]>> + Cases_on`fc`>>fs[mk_vomap_def]>> + qmatch_asmsub_abbrev_tac`strlit A`>> + qmatch_goalsub_abbrev_tac`strlit B`>> + qsuff_tac`A = B`>- metis_tac[]>> + unabbrev_all_tac>> + simp[map_snd_def]>> + match_mp_tac map_foldl_rel>> + simp[map_replicate]>> + EVAL_TAC +QED val check_cstep_arr = process_topdecs` fun check_cstep_arr lno cstep fml inds vimap vomap pc = @@ -3352,7 +3458,7 @@ val check_cstep_arr = process_topdecs` (case check_change_obj_arr lno b fml (get_id pc) (get_obj pc) fc' pfs of (fml',(fc',id')) => - (fml', (inds, (vimap, (mk_vomap fc', change_obj_update pc id' fc'))))) + (fml', (inds, (vimap, (mk_vomap_arr (String.size vomap) fc', change_obj_update pc id' fc'))))) | Checkobj fc' => if check_eq_obj (get_obj pc) fc' then (fml, (inds, (vimap, (vomap, pc)))) diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 583ec68f78..9b2c766234 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -774,7 +774,10 @@ Definition set_indices_def: (inds, sptree$insert n rinds vimap) End -(* Fast substitution for obj_constraint if it is *) +val ow = rconc (EVAL``CHR 1``); +val zw = rconc (EVAL``CHR 0``); + +(* Fast substitution for obj_constraint if it is in vomap *) Definition fast_obj_constraint_def: fast_obj_constraint s l vomap = case s of @@ -782,9 +785,13 @@ Definition fast_obj_constraint_def: if length v = 0 then ([],0) else obj_constraint (subst_fun s) l | INL (n,_) => - case sptree$lookup n vomap of - NONE => ([],0) - | SOME () => obj_constraint (subst_fun s) l + if n < strlen vomap + then + if strsub vomap n = ^zw then + ([],0) + else + obj_constraint (subst_fun s) l + else ([],0) End Definition fast_red_subgoals_def: @@ -1557,11 +1564,13 @@ Proof QED Definition vomap_rel_def: - vomap_rel obj vomap ⇔ + vomap_rel obj ls ⇔ case obj of NONE => T | SOME l => - set (MAP SND (FST l)) = domain vomap + ∀x. + MEM x (MAP SND (FST l)) <=> + x < strlen ls ∧ strsub ls x ≠ ^zw End Theorem add_lists_map_negate_coeff: @@ -2223,15 +2232,32 @@ Definition check_change_obj_list_def: End Definition mk_vomap_def: - mk_vomap (f,c) = - list_to_num_set (MAP SND f) + mk_vomap n (f,c) = + strlit (FOLDL (λacc i. update_resize acc ^zw ^ow i) (REPLICATE n ^zw) (MAP SND f)) End +Theorem resize_acc_bitset_iff: + ∀ls acc. + (x < + LENGTH + (FOLDL (λacc i. update_resize acc ^zw ^ow i) acc ls) ∧ + EL x (FOLDL (λacc i. update_resize acc ^zw ^ow i) acc ls) ≠ ^zw) ⇔ + (MEM x ls ∨ x < LENGTH acc ∧ EL x acc ≠ ^zw) +Proof + Induct>>rw[]>> + rw[update_resize_def,EL_LUPDATE,EL_APPEND_EQN,EL_REPLICATE]>> + EVERY_CASE_TAC>>gvs[]>> + Cases_on`x < 2 * h + 1` >>simp[]>> + DEP_REWRITE_TAC[EL_REPLICATE]>> + simp[] +QED + Theorem vomap_rel_mk_vomap: - vomap_rel (SOME fc) (mk_vomap fc) + vomap_rel (SOME fc) (mk_vomap n fc) Proof Cases_on`fc`>>rw[vomap_rel_def,mk_vomap_def]>> - fs[EXTENSION,domain_list_to_num_set] + simp[resize_acc_bitset_iff]>> + metis_tac[EL_REPLICATE] QED Definition check_cstep_list_def: @@ -2359,7 +2385,7 @@ Definition check_cstep_list_def: NONE => NONE | SOME (fml',fc',id') => SOME ( - fml', inds, vimap, mk_vomap fc', + fml', inds, vimap, mk_vomap (strlen vomap) fc', pc with <| id:=id'; obj:=SOME fc' |>)) | CheckObj fc' => if check_eq_obj pc.obj fc' @@ -2990,8 +3016,8 @@ Proof QED Definition mk_vomap_opt_def: - (mk_vomap_opt NONE = LN) ∧ - (mk_vomap_opt (SOME fc) = mk_vomap fc) + (mk_vomap_opt NONE = strlit "") ∧ + (mk_vomap_opt (SOME fc) = mk_vomap (LENGTH (FST fc)) fc) End Theorem check_csteps_list_concl: @@ -3024,7 +3050,7 @@ Proof unabbrev_all_tac>> simp[init_conf_def]>> Cases_on`obj`>- - simp[vomap_rel_def]>> + EVAL_TAC>> metis_tac[vomap_rel_mk_vomap,mk_vomap_opt_def])>> `∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE` by ( rw[Abbr`pc`,Abbr`fmlls`,any_el_ALT,init_conf_def]>> @@ -3145,7 +3171,7 @@ Proof unabbrev_all_tac>> simp[init_conf_def]>> Cases_on`obj`>- - simp[vomap_rel_def]>> + EVAL_TAC>> metis_tac[vomap_rel_mk_vomap,mk_vomap_opt_def])>> `∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE` by ( rw[Abbr`pc`,Abbr`fmlls`,any_el_ALT,init_conf_def]>> diff --git a/examples/pseudo_bool/array/npbc_parseProgScript.sml b/examples/pseudo_bool/array/npbc_parseProgScript.sml index e604eff267..04771206f1 100644 --- a/examples/pseudo_bool/array/npbc_parseProgScript.sml +++ b/examples/pseudo_bool/array/npbc_parseProgScript.sml @@ -2928,7 +2928,30 @@ val res = translate init_conf_def; val res = translate hconcl_concl_def; val res = translate conv_boutput_hconcl_def; -val res = translate npbc_listTheory.mk_vomap_opt_def; +val mk_vomap_opt_arr = process_topdecs` + fun mk_vomap_opt_arr obj = + case obj of None => "" + | Some fc => + mk_vomap_arr (List.length (fst fc)) fc` |> append_prog; + +Theorem mk_vomap_opt_arr_spec: + obj_TYPE obj objv + ⇒ + app (p : 'ffi ffi_proj) + ^(fetch_v "mk_vomap_opt_arr" (get_ml_prog_state())) + [objv] + (emp) + (POSTv v. &(vomap_TYPE (mk_vomap_opt obj) v)) +Proof + rw[]>> + xcf "mk_vomap_opt_arr" (get_ml_prog_state ())>> + Cases_on`obj`>>fs[OPTION_TYPE_def,npbc_listTheory.mk_vomap_opt_def]>> + xmatch + >- ( + xlit>>xsimpl)>> + rpt xlet_autop>> + xapp>>xsimpl +QED val check_unsat' = process_topdecs ` fun check_unsat' fns fd lno fml obj fmlt objt = @@ -2938,7 +2961,7 @@ val check_unsat' = process_topdecs ` val arr = fill_arr arr 1 fml val inds = rev_enum_full 1 fml val vimap = fold_update_vimap_enum_full 1 fml - val vomap = mk_vomap_opt obj + val vomap = mk_vomap_opt_arr obj val pc = init_conf id True obj in (case check_unsat'' fns fd lno arr inds vimap vomap pc of From ebed25ecea4bb2d1ddd2a06df2515b0279224062 Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sat, 3 Feb 2024 00:42:20 +0800 Subject: [PATCH 22/38] try to speedup fast_obj_constraint --- .../pseudo_bool/array/npbc_listScript.sml | 74 +++++++++++++++---- 1 file changed, 58 insertions(+), 16 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 9b2c766234..ea1dff3b1e 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -774,9 +774,17 @@ Definition set_indices_def: (inds, sptree$insert n rinds vimap) End +(* ow indicates positive literal occurence + tw indicates negative literal occurence *) val ow = rconc (EVAL``CHR 1``); +val tw = rconc (EVAL``CHR 2``); val zw = rconc (EVAL``CHR 0``); +Definition match_polarity_def: + match_polarity w b ⇔ + (w = ^ow ∧ ¬ b) ∨ (w = ^tw ∧ b) +End + (* Fast substitution for obj_constraint if it is in vomap *) Definition fast_obj_constraint_def: fast_obj_constraint s l vomap = @@ -784,13 +792,18 @@ Definition fast_obj_constraint_def: INR v => if length v = 0 then ([],0) else obj_constraint (subst_fun s) l - | INL (n,_) => + | INL (n,v) => if n < strlen vomap then if strsub vomap n = ^zw then ([],0) else - obj_constraint (subst_fun s) l + case v of + INL b => + if match_polarity (strsub vomap n) b + then ([],0) + else obj_constraint (subst_fun s) l + | _ => obj_constraint (subst_fun s) l else ([],0) End @@ -1568,9 +1581,12 @@ Definition vomap_rel_def: case obj of NONE => T | SOME l => - ∀x. - MEM x (MAP SND (FST l)) <=> - x < strlen ls ∧ strsub ls x ≠ ^zw + ∀c x. + MEM (c:int,x) (FST l) <=> + x < strlen ls ∧ + strsub ls x ≠ ^zw ∧ + (strsub ls x = ^ow ⇒ c > 0) ∧ + (strsub ls x = ^tw ⇒ c < 0) End Theorem add_lists_map_negate_coeff: @@ -1610,19 +1626,45 @@ Theorem vomap_rel_fast_obj_constraint: obj_constraint (subst_fun s) l Proof rw[fast_obj_constraint_def]>> - every_case_tac>> Cases_on`l`>> fs[npbcTheory.obj_constraint_def,subst_fun_def]>> - rpt (pairarg_tac>>fs[])>> - pop_assum mp_tac>> - DEP_REWRITE_TAC[add_lists_map_negate_coeff]>>rw[]>> - pop_assum mp_tac>> - DEP_REWRITE_TAC[subst_lhs_id]>> - fs[vomap_rel_def]>> - simp[EVERY_MAP,LAMBDA_PROD,subst_fun_def]>> - gvs[EVERY_MEM]>> - rw[]>>pairarg_tac>>fs[EXTENSION,MEM_MAP,domain_lookup]>> - metis_tac[option_CLAUSES,SND,PAIR] + reverse TOP_CASE_TAC>>gvs[] + >- ( + (* INR *) + rw[]>> + rpt (pairarg_tac>>fs[])>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[add_lists_map_negate_coeff]>>rw[]>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[subst_lhs_id]>> + simp[EVERY_MAP,subst_fun_def,EVERY_MEM])>> + TOP_CASE_TAC>>simp[]>> + reverse TOP_CASE_TAC>>simp[] + >- ( + (* not in map *) + rw[]>> + rpt (pairarg_tac>>fs[])>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[add_lists_map_negate_coeff]>>rw[]>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[subst_lhs_id]>> + simp[EVERY_MAP,subst_fun_def,EVERY_MEM]>> + fs[vomap_rel_def,FORALL_PROD]) >> + TOP_CASE_TAC>>simp[] + >- ( + (* not in map *) + rw[]>> + rpt (pairarg_tac>>fs[])>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[add_lists_map_negate_coeff]>>rw[]>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[subst_lhs_id]>> + simp[EVERY_MAP,subst_fun_def,EVERY_MEM]>> + fs[vomap_rel_def,FORALL_PROD])>> + TOP_CASE_TAC>>simp[]>> + TOP_CASE_TAC>>simp[]>> + rpt(pairarg_tac>>gvs[])>> + cheat QED Theorem vomap_rel_fast_red_subgoals: From 3210db444c0b32184a3841f562a9aa5eeec7bcab Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sat, 3 Feb 2024 18:13:20 +0800 Subject: [PATCH 23/38] Revert "try to speedup fast_obj_constraint" This reverts commit ebed25ecea4bb2d1ddd2a06df2515b0279224062. --- .../pseudo_bool/array/npbc_listScript.sml | 74 ++++--------------- 1 file changed, 16 insertions(+), 58 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index ea1dff3b1e..9b2c766234 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -774,17 +774,9 @@ Definition set_indices_def: (inds, sptree$insert n rinds vimap) End -(* ow indicates positive literal occurence - tw indicates negative literal occurence *) val ow = rconc (EVAL``CHR 1``); -val tw = rconc (EVAL``CHR 2``); val zw = rconc (EVAL``CHR 0``); -Definition match_polarity_def: - match_polarity w b ⇔ - (w = ^ow ∧ ¬ b) ∨ (w = ^tw ∧ b) -End - (* Fast substitution for obj_constraint if it is in vomap *) Definition fast_obj_constraint_def: fast_obj_constraint s l vomap = @@ -792,18 +784,13 @@ Definition fast_obj_constraint_def: INR v => if length v = 0 then ([],0) else obj_constraint (subst_fun s) l - | INL (n,v) => + | INL (n,_) => if n < strlen vomap then if strsub vomap n = ^zw then ([],0) else - case v of - INL b => - if match_polarity (strsub vomap n) b - then ([],0) - else obj_constraint (subst_fun s) l - | _ => obj_constraint (subst_fun s) l + obj_constraint (subst_fun s) l else ([],0) End @@ -1581,12 +1568,9 @@ Definition vomap_rel_def: case obj of NONE => T | SOME l => - ∀c x. - MEM (c:int,x) (FST l) <=> - x < strlen ls ∧ - strsub ls x ≠ ^zw ∧ - (strsub ls x = ^ow ⇒ c > 0) ∧ - (strsub ls x = ^tw ⇒ c < 0) + ∀x. + MEM x (MAP SND (FST l)) <=> + x < strlen ls ∧ strsub ls x ≠ ^zw End Theorem add_lists_map_negate_coeff: @@ -1626,45 +1610,19 @@ Theorem vomap_rel_fast_obj_constraint: obj_constraint (subst_fun s) l Proof rw[fast_obj_constraint_def]>> + every_case_tac>> Cases_on`l`>> fs[npbcTheory.obj_constraint_def,subst_fun_def]>> - reverse TOP_CASE_TAC>>gvs[] - >- ( - (* INR *) - rw[]>> - rpt (pairarg_tac>>fs[])>> - pop_assum mp_tac>> - DEP_REWRITE_TAC[add_lists_map_negate_coeff]>>rw[]>> - pop_assum mp_tac>> - DEP_REWRITE_TAC[subst_lhs_id]>> - simp[EVERY_MAP,subst_fun_def,EVERY_MEM])>> - TOP_CASE_TAC>>simp[]>> - reverse TOP_CASE_TAC>>simp[] - >- ( - (* not in map *) - rw[]>> - rpt (pairarg_tac>>fs[])>> - pop_assum mp_tac>> - DEP_REWRITE_TAC[add_lists_map_negate_coeff]>>rw[]>> - pop_assum mp_tac>> - DEP_REWRITE_TAC[subst_lhs_id]>> - simp[EVERY_MAP,subst_fun_def,EVERY_MEM]>> - fs[vomap_rel_def,FORALL_PROD]) >> - TOP_CASE_TAC>>simp[] - >- ( - (* not in map *) - rw[]>> - rpt (pairarg_tac>>fs[])>> - pop_assum mp_tac>> - DEP_REWRITE_TAC[add_lists_map_negate_coeff]>>rw[]>> - pop_assum mp_tac>> - DEP_REWRITE_TAC[subst_lhs_id]>> - simp[EVERY_MAP,subst_fun_def,EVERY_MEM]>> - fs[vomap_rel_def,FORALL_PROD])>> - TOP_CASE_TAC>>simp[]>> - TOP_CASE_TAC>>simp[]>> - rpt(pairarg_tac>>gvs[])>> - cheat + rpt (pairarg_tac>>fs[])>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[add_lists_map_negate_coeff]>>rw[]>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[subst_lhs_id]>> + fs[vomap_rel_def]>> + simp[EVERY_MAP,LAMBDA_PROD,subst_fun_def]>> + gvs[EVERY_MEM]>> + rw[]>>pairarg_tac>>fs[EXTENSION,MEM_MAP,domain_lookup]>> + metis_tac[option_CLAUSES,SND,PAIR] QED Theorem vomap_rel_fast_red_subgoals: From c210cf6122412ff04981191d65cc61d4ea0b6e5e Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sat, 3 Feb 2024 19:37:33 +0800 Subject: [PATCH 24/38] make npbc_list build --- .../pseudo_bool/array/npbc_listScript.sml | 279 ++++++++++++------ 1 file changed, 192 insertions(+), 87 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 9b2c766234..33b0e49b71 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -774,8 +774,28 @@ Definition set_indices_def: (inds, sptree$insert n rinds vimap) End -val ow = rconc (EVAL``CHR 1``); -val zw = rconc (EVAL``CHR 0``); +Definition byte_strsub_def: + byte_strsub ls x = + (n2w (ORD (strsub ls x))):word8 +End + +Definition vomap_rel_def: + vomap_rel obj ls ⇔ + case obj of + NONE => T + | SOME (l,n) => + ∀x. + (x < strlen ls ⇒ + let w = byte_strsub ls x in + if w = 0w then ¬MEM x (MAP SND l) + else + w ≠ 128w ⇒ + ∃i. + i < LENGTH l ∧ + FST (EL i l) = w2i w ∧ + SND (EL i l) = x ) ∧ + (~(x < strlen ls) ⇒ ¬MEM x (MAP SND l)) +End (* Fast substitution for obj_constraint if it is in vomap *) Definition fast_obj_constraint_def: @@ -784,16 +804,140 @@ Definition fast_obj_constraint_def: INR v => if length v = 0 then ([],0) else obj_constraint (subst_fun s) l - | INL (n,_) => + | INL (n,v) => if n < strlen vomap then - if strsub vomap n = ^zw then - ([],0) + let w = byte_strsub vomap n in + if w = 0w then ([],0) + else if w = 128w then obj_constraint (subst_fun s) l else - obj_constraint (subst_fun s) l + case v of INL b => + let c = w2i w in + ([c, n], if 0 ≤ c ⇔ b then Num (ABS c) else 0) + | _ => obj_constraint (subst_fun s) l else ([],0) End +Theorem subst_aux_id: + ∀l. + EVERY (\v. f v = NONE) (MAP SND l) ⇒ + subst_aux f l = (l,[],0) +Proof + Induct>-simp[npbcTheory.subst_aux_def]>> + Cases>> + rw[npbcTheory.subst_aux_def] +QED + +Theorem subst_lhs_id: + EVERY (\v. f v = NONE) (MAP SND l) ⇒ + subst_lhs f l = (l, 0) +Proof + rw[npbcTheory.subst_lhs_def]>> + rpt(pairarg_tac>>fs[])>> + drule subst_aux_id>>strip_tac>> + gvs[EVAL``clean_up []``]>> + Cases_on`l`>> + gvs[npbcTheory.add_lists_def] +QED + +Theorem subst_lhs_subst_fun_bool: + ∀l. + subst_aux (subst_fun (INL (x,INL b))) l = + (FILTER (λ(c,y). x ≠ y) l, [], + SUM (MAP (Num o ABS o FST) (FILTER (λ(c,y). x = y ∧ (0 ≤ c ⇔ b)) l)) + ) +Proof + Induct>>rw[npbcTheory.subst_aux_def]>> + pairarg_tac>>gvs[npbcTheory.subst_aux_def]>> + simp[subst_fun_def] +QED + +Theorem add_lists_emp_2: + add_lists ls [] = (ls,0) +Proof + Cases_on`ls`>>EVAL_TAC +QED + +Theorem add_lists_map_negate_coeff: + ∀ls rs. + rs = (MAP (λ(c,l). (-c,l)) ls) ⇒ + add_lists ls rs = ([],SUM (MAP (λi. Num (ABS (FST i))) ls)) +Proof + ho_match_mp_tac npbcTheory.add_lists_ind>> + simp[npbcTheory.add_lists_def,npbcTheory.add_terms_def] +QED + +Theorem add_lists_FILTER_map_negate_coeff: + ∀ls rs. + SORTED $< (MAP SND ls) ∧ + i < LENGTH ls ∧ + FILTER (λ(c,n). SND (EL i ls) ≠ n) (MAP (λ(c,l). (-c,l)) ls) = rs ⇒ + add_lists ls rs = + ([EL i ls], + SUM (MAP (λi. Num (ABS (FST i))) ls) - + Num (ABS (FST (EL i ls)))) +Proof + cheat +QED + +Theorem vomap_rel_fast_obj_constraint: + SORTED $< (MAP SND (FST l)) ∧ + vomap_rel (SOME l) vomap ⇒ + fast_obj_constraint s l vomap = + obj_constraint (subst_fun s) l +Proof + rw[fast_obj_constraint_def]>> + Cases_on`l`>> + fs[npbcTheory.obj_constraint_def,subst_fun_def]>> + reverse TOP_CASE_TAC>>gvs[] + >- ( + (* INR *) + rw[]>> + rpt (pairarg_tac>>fs[])>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[add_lists_map_negate_coeff]>>rw[]>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[subst_lhs_id]>> + simp[EVERY_MAP,subst_fun_def,EVERY_MEM])>> + TOP_CASE_TAC>>simp[]>> + fs[vomap_rel_def,FORALL_PROD,FORALL_AND_THM]>> + reverse TOP_CASE_TAC>>simp[] + >- ( + (* not in map *) + first_x_assum drule>> + rw[]>> + rpt (pairarg_tac>>fs[])>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[add_lists_map_negate_coeff]>>rw[]>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[subst_lhs_id]>> + simp[EVERY_MAP,subst_fun_def,EVERY_MEM]>> + simp[FORALL_PROD]>> + metis_tac[PAIR,SND,MEM_MAP])>> + first_x_assum drule>>rw[]>>fs[] + >- ( + (* not in map *) + rpt (pairarg_tac>>fs[])>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[add_lists_map_negate_coeff]>>rw[]>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[subst_lhs_id]>> + simp[EVERY_MAP,subst_fun_def,EVERY_MEM,FORALL_PROD]>> + rw[]>>metis_tac[PAIR,SND,MEM_MAP])>> + TOP_CASE_TAC>>simp[]>> + simp[npbcTheory.subst_lhs_def]>> + simp[subst_lhs_subst_fun_bool,npbcTheory.clean_up_def,add_lists_emp_2]>> + drule add_lists_FILTER_map_negate_coeff>> + disch_then drule>> + disch_then (fn th => DEP_REWRITE_TAC[GEN_ALL th])>> + simp[]>> + CONJ_TAC >- + metis_tac[PAIR]>> + qmatch_goalsub_abbrev_tac`A - ( _ - B + C)`>> + qmatch_goalsub_abbrev_tac`w2i bb`>> + cheat +QED + Definition fast_red_subgoals_def: fast_red_subgoals ord s def obj vomap = let cobj = @@ -1563,75 +1707,20 @@ Proof simp[any_el_ALT] QED -Definition vomap_rel_def: - vomap_rel obj ls ⇔ - case obj of - NONE => T - | SOME l => - ∀x. - MEM x (MAP SND (FST l)) <=> - x < strlen ls ∧ strsub ls x ≠ ^zw +Definition sorted_obj_def: + sorted_obj obj ⇔ + case obj of NONE => T + | SOME (l:(int # num) list # int) => SORTED $< (MAP SND (FST l)) End -Theorem add_lists_map_negate_coeff: - ∀ls rs. - rs = (MAP (λ(c,l). (-c,l)) ls) ⇒ - add_lists ls rs = ([],SUM (MAP (λi. Num (ABS (FST i))) ls)) -Proof - ho_match_mp_tac npbcTheory.add_lists_ind>> - simp[npbcTheory.add_lists_def,npbcTheory.add_terms_def] -QED - -Theorem subst_aux_id: - ∀l. - EVERY (\v. f v = NONE) (MAP SND l) ⇒ - subst_aux f l = (l,[],0) -Proof - Induct>-simp[npbcTheory.subst_aux_def]>> - Cases>> - rw[npbcTheory.subst_aux_def] -QED - -Theorem subst_lhs_id: - EVERY (\v. f v = NONE) (MAP SND l) ⇒ - subst_lhs f l = (l, 0) -Proof - rw[npbcTheory.subst_lhs_def]>> - rpt(pairarg_tac>>fs[])>> - drule subst_aux_id>>strip_tac>> - gvs[EVAL``clean_up []``]>> - Cases_on`l`>> - gvs[npbcTheory.add_lists_def] -QED - -Theorem vomap_rel_fast_obj_constraint: - vomap_rel (SOME l) vomap ⇒ - fast_obj_constraint s l vomap = - obj_constraint (subst_fun s) l -Proof - rw[fast_obj_constraint_def]>> - every_case_tac>> - Cases_on`l`>> - fs[npbcTheory.obj_constraint_def,subst_fun_def]>> - rpt (pairarg_tac>>fs[])>> - pop_assum mp_tac>> - DEP_REWRITE_TAC[add_lists_map_negate_coeff]>>rw[]>> - pop_assum mp_tac>> - DEP_REWRITE_TAC[subst_lhs_id]>> - fs[vomap_rel_def]>> - simp[EVERY_MAP,LAMBDA_PROD,subst_fun_def]>> - gvs[EVERY_MEM]>> - rw[]>>pairarg_tac>>fs[EXTENSION,MEM_MAP,domain_lookup]>> - metis_tac[option_CLAUSES,SND,PAIR] -QED - Theorem vomap_rel_fast_red_subgoals: + sorted_obj obj ∧ vomap_rel obj vomap ⇒ fast_red_subgoals ord s def obj vomap = red_subgoals ord (subst_fun s) def obj Proof rw[fast_red_subgoals_def,red_subgoals_def]>> - every_case_tac>>fs[]>> + every_case_tac>>fs[sorted_obj_def]>> metis_tac[vomap_rel_fast_obj_constraint] QED @@ -1639,6 +1728,7 @@ Theorem fml_rel_check_red_list: fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ vimap_rel fmlls vimap ∧ + sorted_obj obj ∧ vomap_rel obj vomap ∧ (∀n. n ≥ id ⇒ any_el n fmlls NONE = NONE) ∧ check_red_list ord obj b tcb fmlls inds id c s pfs @@ -2064,6 +2154,7 @@ Theorem fml_rel_check_sstep_list: fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ vimap_rel fmlls vimap ∧ + sorted_obj obj ∧ vomap_rel obj vomap ∧ (∀n. n ≥ id ⇒ any_el n fmlls NONE = NONE) ∧ check_sstep_list sstep ord obj tcb fmlls inds id vimap vomap = @@ -2231,33 +2322,32 @@ Definition check_change_obj_list_def: else NONE) End +Definition upd_vomap_def: + upd_vomap c n = + if c = 0 ∨ ABS c ≥ 128 then + CHR 128 + else + let (w:word8) = i2w c in + CHR (w2n w) +End + Definition mk_vomap_def: mk_vomap n (f,c) = - strlit (FOLDL (λacc i. update_resize acc ^zw ^ow i) (REPLICATE n ^zw) (MAP SND f)) + strlit (FOLDL (λacc (c,n). + update_resize acc (CHR 0) (upd_vomap c n) n ) + (REPLICATE n (CHR 0)) f) End -Theorem resize_acc_bitset_iff: - ∀ls acc. - (x < - LENGTH - (FOLDL (λacc i. update_resize acc ^zw ^ow i) acc ls) ∧ - EL x (FOLDL (λacc i. update_resize acc ^zw ^ow i) acc ls) ≠ ^zw) ⇔ - (MEM x ls ∨ x < LENGTH acc ∧ EL x acc ≠ ^zw) -Proof - Induct>>rw[]>> - rw[update_resize_def,EL_LUPDATE,EL_APPEND_EQN,EL_REPLICATE]>> - EVERY_CASE_TAC>>gvs[]>> - Cases_on`x < 2 * h + 1` >>simp[]>> - DEP_REWRITE_TAC[EL_REPLICATE]>> - simp[] -QED - Theorem vomap_rel_mk_vomap: + SORTED $< (MAP SND (FST fc)) ⇒ vomap_rel (SOME fc) (mk_vomap n fc) Proof Cases_on`fc`>>rw[vomap_rel_def,mk_vomap_def]>> + rw[]>> + cheat + (* simp[resize_acc_bitset_iff]>> - metis_tac[EL_REPLICATE] + metis_tac[EL_REPLICATE] *) QED Definition check_cstep_list_def: @@ -2536,6 +2626,7 @@ Theorem fml_rel_check_cstep_list: fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ vimap_rel fmlls vimap ∧ + sorted_obj pc.obj ∧ vomap_rel pc.obj vomap ∧ (∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE) ∧ check_cstep_list cstep fmlls inds vimap vomap pc = @@ -2545,6 +2636,7 @@ Theorem fml_rel_check_cstep_list: fml_rel fml' fmlls' ∧ ind_rel fmlls' inds' ∧ vimap_rel fmlls' vimap' ∧ + sorted_obj pc'.obj ∧ vomap_rel pc'.obj vomap' ∧ (∀n. n ≥ pc'.id ⇒ any_el n fmlls' NONE = NONE) ∧ pc.id ≤ pc'.id @@ -2751,8 +2843,11 @@ Proof metis_tac[ind_rel_reindex])>> CONJ_TAC >- metis_tac[fml_rel_fml_rel_vimap_rel]>> - CONJ_TAC >- - metis_tac[vomap_rel_mk_vomap]>> + CONJ_ASM1_TAC >- + cheat>> + CONJ_TAC >- ( + fs[sorted_obj_def]>> + metis_tac[vomap_rel_mk_vomap])>> simp[any_el_rollback]) >- ( (* CheckObj *) fs[check_cstep_def,check_cstep_list_def] @@ -2775,6 +2870,7 @@ Theorem fml_rel_check_csteps_list: fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ vimap_rel fmlls vimap ∧ + sorted_obj pc.obj ∧ vomap_rel pc.obj vomap ∧ (∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE) ∧ check_csteps_list csteps fmlls inds vimap vomap pc = @@ -2784,6 +2880,7 @@ Theorem fml_rel_check_csteps_list: fml_rel fml' fmlls' ∧ ind_rel fmlls' inds' ∧ vimap_rel fmlls' vimap' ∧ + sorted_obj pc'.obj ∧ vomap_rel pc'.obj vomap' ∧ (∀n. n ≥ pc'.id ⇒ any_el n fmlls' NONE = NONE) ∧ pc.id ≤ pc'.id @@ -3021,6 +3118,7 @@ Definition mk_vomap_opt_def: End Theorem check_csteps_list_concl: + sorted_obj obj ∧ check_csteps_list cs (FOLDL (λacc (i,v). update_resize acc NONE (SOME (v,T)) i) (REPLICATE m NONE) (enumerate 1 fml)) @@ -3046,11 +3144,14 @@ Proof `vimap_rel fmlls vimap` by ( unabbrev_all_tac>> simp[vimap_rel_FOLDL_update_resize])>> + `sorted_obj pc.obj` by + gvs[sorted_obj_def,init_conf_def,Abbr`pc`]>> `vomap_rel pc.obj vomap` by ( unabbrev_all_tac>> simp[init_conf_def]>> Cases_on`obj`>- EVAL_TAC>> + gvs[sorted_obj_def]>> metis_tac[vomap_rel_mk_vomap,mk_vomap_opt_def])>> `∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE` by ( rw[Abbr`pc`,Abbr`fmlls`,any_el_ALT,init_conf_def]>> @@ -3143,6 +3244,7 @@ Proof QED Theorem check_csteps_list_output: + sorted_obj obj ∧ check_csteps_list cs (FOLDL (λacc (i,v). update_resize acc NONE (SOME (v,T)) i) (REPLICATE m NONE) (enumerate 1 fml)) @@ -3167,11 +3269,14 @@ Proof `vimap_rel fmlls vimap` by ( unabbrev_all_tac>> simp[vimap_rel_FOLDL_update_resize])>> + `sorted_obj pc.obj` by + gvs[sorted_obj_def,init_conf_def,Abbr`pc`]>> `vomap_rel pc.obj vomap` by ( unabbrev_all_tac>> simp[init_conf_def]>> Cases_on`obj`>- EVAL_TAC>> + gvs[sorted_obj_def]>> metis_tac[vomap_rel_mk_vomap,mk_vomap_opt_def])>> `∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE` by ( rw[Abbr`pc`,Abbr`fmlls`,any_el_ALT,init_conf_def]>> From a9a9983b1110ad892471da6913c45fae1ce43f73 Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sat, 3 Feb 2024 20:08:48 +0800 Subject: [PATCH 25/38] make translation work --- .../array/npbc_arrayProgScript.sml | 58 ++++++++++--------- .../pseudo_bool/array/npbc_listScript.sml | 12 ++-- 2 files changed, 38 insertions(+), 32 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_arrayProgScript.sml b/examples/pseudo_bool/array/npbc_arrayProgScript.sml index a4d551b2bc..5624a0e12c 100644 --- a/examples/pseudo_bool/array/npbc_arrayProgScript.sml +++ b/examples/pseudo_bool/array/npbc_arrayProgScript.sml @@ -1554,7 +1554,14 @@ End val res = translate list_list_insert_def; val res = translate npbcTheory.dom_subst_def; -val res = translate fast_obj_constraint_def; + +val r = translate byte_strsub_def; +val r = translate fast_obj_constraint_def; + +val fast_obj_constraint_side = Q.prove(` + fast_obj_constraint_side x y z ⇔ T`, + EVAL_TAC>>rw[]) |> update_precondition; + val res = translate fast_red_subgoals_def; val res = translate do_rso_def; @@ -3234,84 +3241,84 @@ val res = translate err_obj_check_string_def; val res = translate npbc_checkTheory.eq_obj_def; val res = translate npbc_checkTheory.check_eq_obj_def; +val res = translate upd_vomap_def; + Definition w8z_def: w8z = (0w: word8) End -Definition w8o_def: - w8o = (1w: word8) -End - val w8z_v_thm = translate w8z_def; -val w8o_v_thm = translate w8o_def; val fold_update_resize_bitset = process_topdecs` fun fold_update_resize_bitset ls acc = case ls of [] => acc - | (x::xs) => + | (cn::xs) => case cn of (c,x) => if x < Word8Array.length acc then - (Word8Array.update acc x w8o; + (Word8Array.update acc x (upd_vomap c); fold_update_resize_bitset xs acc) else let val arr = Word8Array.array (2*x+1) w8z val u = Word8Array.copy acc 0 (Word8Array.length acc) arr 0 in - (Word8Array.update arr x w8o; + (Word8Array.update arr x (upd_vomap c); fold_update_resize_bitset xs arr) end ` |> append_prog; Theorem fold_update_resize_bitset_spec: ∀ls lsv accv accls. - LIST_TYPE NUM ls lsv + LIST_TYPE (PAIR_TYPE INT NUM) ls lsv ⇒ app (p : 'ffi ffi_proj) ^(fetch_v "fold_update_resize_bitset" (get_ml_prog_state())) [lsv; accv] (W8ARRAY accv accls) (POSTv v. - W8ARRAY v (FOLDL (λacc i. update_resize acc w8z w8o i) accls ls)) + W8ARRAY v (FOLDL (λacc (c,n). update_resize acc 0w (upd_vomap c) n) accls ls)) Proof Induct>> xcf "fold_update_resize_bitset" (get_ml_prog_state ())>> gvs[LIST_TYPE_def]>>xmatch - >- ( - xvar>>xsimpl)>> - assume_tac w8o_v_thm>> - assume_tac w8z_v_thm>> + >- (xvar>>xsimpl)>> + pairarg_tac>>fs[PAIR_TYPE_def]>> + xmatch>> rpt xlet_autop>> xif >- ( + xlet_autop>> xlet_autop>> xapp>>xsimpl>> simp[update_resize_def])>> + assume_tac w8z_v_thm>> rpt xlet_autop>> xapp>>xsimpl>> - simp[update_resize_def] + simp[update_resize_def,w8z_def] QED val mk_vomap_arr = process_topdecs` fun mk_vomap_arr n fc = let val acc = Word8Array.array n w8z - val f = map_snd (fst fc) + val f = fst fc val acc = fold_update_resize_bitset f acc in Word8Array.substring acc 0 (Word8Array.length acc) end` |> append_prog; Theorem map_foldl_rel: ∀ls accA accB. - MAP (CHR o w2n) accA = accB ⇒ + MAP (CHR ∘ w2n) accA = accB ∧ + CHR (w2n wz) = cz ∧ + (!i. CHR (w2n (wo i)) = co i) ⇒ MAP (CHR ∘ w2n) - (FOLDL (λacc i. update_resize acc w8z w8o i) accA ls) = - FOLDL (λacc i. update_resize acc #"\^@" #"\^A" i) accB ls + (FOLDL (λacc (c,n). update_resize acc wz (wo c) n) accA ls) = + FOLDL (λacc (c,n). update_resize acc cz (co c) n) accB ls Proof Induct>>rw[]>> - first_x_assum match_mp_tac>> - rw[update_resize_def,LUPDATE_MAP]>> - EVAL_TAC + gvs[]>> + pairarg_tac>>rw[]>>gvs[]>> + rw[update_resize_def,LUPDATE_MAP] QED Theorem mk_vomap_arr_spec: @@ -3338,9 +3345,8 @@ Proof qmatch_goalsub_abbrev_tac`strlit B`>> qsuff_tac`A = B`>- metis_tac[]>> unabbrev_all_tac>> - simp[map_snd_def]>> - match_mp_tac map_foldl_rel>> - simp[map_replicate]>> + ho_match_mp_tac map_foldl_rel>> + simp[]>> EVAL_TAC QED diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 33b0e49b71..202298e1ee 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -2323,18 +2323,18 @@ Definition check_change_obj_list_def: End Definition upd_vomap_def: - upd_vomap c n = - if c = 0 ∨ ABS c ≥ 128 then - CHR 128 + upd_vomap c = + if c = 0 ∨ Num (ABS c) ≥ 128 then + (128w : word8) else - let (w:word8) = i2w c in - CHR (w2n w) + i2w c End Definition mk_vomap_def: mk_vomap n (f,c) = strlit (FOLDL (λacc (c,n). - update_resize acc (CHR 0) (upd_vomap c n) n ) + update_resize acc (CHR 0) + (CHR (w2n (upd_vomap c))) n ) (REPLICATE n (CHR 0)) f) End From 330c277822217e695cd4c01a173790eb8191c89c Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sat, 3 Feb 2024 20:56:46 +0800 Subject: [PATCH 26/38] check SORTED before SORTING --- examples/pseudo_bool/array/npbc_parseProgScript.sml | 1 + examples/pseudo_bool/pbc_normaliseScript.sml | 6 ++++-- 2 files changed, 5 insertions(+), 2 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_parseProgScript.sml b/examples/pseudo_bool/array/npbc_parseProgScript.sml index 04771206f1..a8974dfee3 100644 --- a/examples/pseudo_bool/array/npbc_parseProgScript.sml +++ b/examples/pseudo_bool/array/npbc_parseProgScript.sml @@ -2032,6 +2032,7 @@ val res = translate parse_assg_def; val res = translate parse_obj_term_def; val res = translate strip_obju_end_def; +val res = translate SORTED_DEF; val res = translate normalise_obj_def; val res = translate parse_obj_term_npbc_def; diff --git a/examples/pseudo_bool/pbc_normaliseScript.sml b/examples/pseudo_bool/pbc_normaliseScript.sml index 3f4d5f72c2..13a48db089 100644 --- a/examples/pseudo_bool/pbc_normaliseScript.sml +++ b/examples/pseudo_bool/pbc_normaliseScript.sml @@ -1125,7 +1125,9 @@ QED Definition normalise_obj_def: (normalise_obj NONE = NONE) ∧ (normalise_obj (SOME (f,c)) = - let (f',c') = compact_lhs (QSORT term_le f) 0 in + let f = + (if SORTED term_le f then f else QSORT term_le f) in + let (f',c') = compact_lhs f 0 in let (f'', c'') = normalise_lhs f' [] 0 in SOME (f'',c + c'+c'')) End @@ -1150,7 +1152,7 @@ Proof disch_then(qspec_then`w` assume_tac)>> drule compact_lhs_sound>> disch_then(qspec_then`w` assume_tac)>> - fs[]>> + every_case_tac>>fs[]>> intLib.ARITH_TAC QED From 8d914b3575895c50ae38d8a0065376b3a807745a Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sat, 3 Feb 2024 20:57:11 +0800 Subject: [PATCH 27/38] fix names in WCNF --- examples/pseudo_bool/array/wcnfProgScript.sml | 32 ++--- .../cnf_encoding/wcnf_to_pbScript.sml | 136 ++++++++++-------- 2 files changed, 91 insertions(+), 77 deletions(-) diff --git a/examples/pseudo_bool/array/wcnfProgScript.sml b/examples/pseudo_bool/array/wcnfProgScript.sml index 4d6f59aeb3..affa021e59 100644 --- a/examples/pseudo_bool/array/wcnfProgScript.sml +++ b/examples/pseudo_bool/array/wcnfProgScript.sml @@ -372,16 +372,16 @@ Definition maxsat_sem_def: case copt of NONE => T | SOME (lbg,ubg) => (case lbg of - NONE => min_unsat wfml = NONE + NONE => opt_cost wfml = NONE | SOME lb => (case ubg of - NONE => (∀w. satisfies_hard w wfml ⇒ lb ≤ weight w wfml) + NONE => (∀w. sat_hard w wfml ⇒ lb ≤ cost w wfml) | SOME ub => if lb = ub then - min_unsat wfml = SOME lb + opt_cost wfml = SOME lb else - (∀w. satisfies_hard w wfml ⇒ lb ≤ weight w wfml) ∧ - (∃w. satisfies_hard w wfml ∧ weight w wfml ≤ ub))) + (∀w. sat_hard w wfml ⇒ lb ≤ cost w wfml) ∧ + (∃w. sat_hard w wfml ∧ cost w wfml ≤ ub))) End Definition print_maxsat_str_def: @@ -394,15 +394,15 @@ Definition print_maxsat_str_def: (case lbg of NONE => strlit "s VERIFIED BOUNDS " ^ - strlit "sum(unsat weights) <= " ^ toString ub ^ strlit"\n" + strlit "COST <= " ^ toString ub ^ strlit"\n" | SOME lb => if lb = ub then - strlit "s VERIFIED MIN sum(unsat weights) = " ^ + strlit "s VERIFIED OPTIMAL COST = " ^ toString ub ^ strlit"\n" else strlit "s VERIFIED " ^ (toString lb) ^ - strlit " <= MIN sum(unsat weights) <= " ^ toString ub ^ strlit"\n")) + strlit " <= COST <= " ^ toString ub ^ strlit"\n")) End Definition check_unsat_2_sem_def: @@ -533,17 +533,17 @@ Proof qexists_tac`x`>>simp[maxsat_sem_def]>> every_case_tac>>fs[] >- ( - (drule_at Any) full_encode_sem_concl_min_unsat>> + (drule_at Any) full_encode_sem_concl_opt_cost>> metis_tac[PAIR]) >- ( (drule_at Any) full_encode_sem_concl>> fs[]>> disch_then (drule_at Any)>>simp[]) >- ( - (drule_at Any) full_encode_sem_concl_min_unsat>> + (drule_at Any) full_encode_sem_concl_opt_cost>> metis_tac[PAIR]) >- ( - (drule_at Any) full_encode_sem_concl_min_unsat>> + (drule_at Any) full_encode_sem_concl_opt_cost>> metis_tac[PAIR]) >- ( (drule_at Any) full_encode_sem_concl>> @@ -604,7 +604,7 @@ QED Definition maxsat_output_sem_def: maxsat_output_sem wfml wfml' iseqopt ⇔ - (iseqopt ⇒ min_unsat wfml = min_unsat wfml') + (iseqopt ⇒ opt_cost wfml = opt_cost wfml') End Definition print_maxsat_output_str_def: @@ -764,24 +764,24 @@ Proof fs[maxsat_sem_def]>> every_case_tac>>fs[] >- ( - (drule_at Any) full_encode_sem_concl_min_unsat>> + (drule_at Any) full_encode_sem_concl_opt_cost>> metis_tac[PAIR]) >- ( (drule_at Any) full_encode_sem_concl>> fs[]>> disch_then (drule_at Any)>>simp[]) >- ( - (drule_at Any) full_encode_sem_concl_min_unsat>> + (drule_at Any) full_encode_sem_concl_opt_cost>> metis_tac[PAIR]) >- ( - (drule_at Any) full_encode_sem_concl_min_unsat>> + (drule_at Any) full_encode_sem_concl_opt_cost>> metis_tac[PAIR]) >- ( (drule_at Any) full_encode_sem_concl>> fs[]>> disch_then (drule_at Any)>>simp[]))>> rw[]>>fs[]>> - (drule_at Any) full_encode_sem_output_min_unsat>> + (drule_at Any) full_encode_sem_output_opt_cost>> fs[]>> metis_tac[PAIR] QED diff --git a/examples/pseudo_bool/cnf_encoding/wcnf_to_pbScript.sml b/examples/pseudo_bool/cnf_encoding/wcnf_to_pbScript.sml index 2c3efdbc75..9329d0e5a4 100644 --- a/examples/pseudo_bool/cnf_encoding/wcnf_to_pbScript.sml +++ b/examples/pseudo_bool/cnf_encoding/wcnf_to_pbScript.sml @@ -18,42 +18,56 @@ Type wcclause = ``:num # cclause``; (* Weighted CNFs are a list of weighted soft clauses *) Type wccnf = ``:wcclause list``; +(* A cleaner definition of satisfaction *) +Definition sat_lit_def: + sat_lit w l ⇔ + l ≠ 0 ∧ + let v = Num (ABS l) in + if l > 0 then w v else ¬ w v +End + +Definition sat_clause_def: + sat_clause w C ⇔ + ∃l. l ∈ set C ∧ sat_lit w l +End + +Definition sat_hard_def: + sat_hard w wfml ⇔ + ∀C. (0:num,C) ∈ set wfml ⇒ sat_clause w C +End + (* The weight of a clause with respect to an assignment (0 if satisfied, w otherwise) *) Definition weight_clause_def: weight_clause w ((n,C):wcclause) = - if satisfies_clause w (interp_cclause C) then 0 else n + if sat_clause w C then 0 else n End -(* The weight of a wcnf is the sum of UNSAT clauses *) -Definition weight_def: - weight w (wfml:wccnf) = - SUM (MAP (weight_clause w) wfml) +Definition cost_def: + cost w wfml = SUM (MAP (weight_clause w) wfml) End -(* Satisfaction of all hard clauses under an assignment *) -Definition satisfies_hard_def: - satisfies_hard w (wfml:wccnf) = - ∀wC. wC ∈ set wfml ∧ FST wC = 0 ⇒ - satisfies_clause w (interp_cclause (SND wC)) +Definition opt_cost_def: + opt_cost wfml = + if ¬∃w. sat_hard w wfml then NONE + else SOME (MIN_SET {cost w wfml | w | sat_hard w wfml}) End -(* Unsatisfiability *) -Definition unsatisfiable_hard_def: - unsatisfiable_hard (wfml:wccnf) = - ∀w. ¬satisfies_hard w wfml -End +(* Some alternative definitions *) +Theorem sat_lit_alt: + sat_lit w l ⇔ + l ≠ 0 ∧ satisfies_literal w (interp_lit l) +Proof + EVAL_TAC>>rw[] +QED -(* The minimum weight of any hard-satisfying assignment - NONE represents UNSAT (or a min weight of infinity) -*) -Definition min_unsat_def: - min_unsat (wfml:wccnf) = - if unsatisfiable_hard wfml then NONE - else - SOME (MIN_SET ( - {weight w wfml | w | satisfies_hard w wfml})) -End +Theorem sat_clause_alt: + sat_clause w C ⇔ + satisfies_clause w (interp_cclause C) +Proof + rw[sat_clause_def,lprTheory.interp_cclause_def,satSemTheory.satisfies_clause_def,PULL_EXISTS]>> + metis_tac[sat_lit_alt] +QED (*** STEP 2: Formalise an encoding into PB ***) @@ -235,7 +249,7 @@ Theorem weight_clause_FILTER: weight_clause w (n,C) = weight_clause w (n, FILTER (λl. l ≠ 0) C) Proof - rw[weight_clause_def]>> + rw[weight_clause_def,sat_clause_alt]>> metis_tac[interp_cclause_FILTER] QED @@ -243,7 +257,7 @@ Theorem weight_clause_nub: weight_clause f (q,nub C) = weight_clause f (q,C) Proof - rw[weight_clause_def,interp_cclause_def] + rw[weight_clause_def,interp_cclause_def,sat_clause_alt] QED Theorem interp_cclause_nub: @@ -278,7 +292,7 @@ Proof `weight_clause (λx. w (INL x)) (q,r) = weight_clause (λx. w (INL x)) (q,C)` by metis_tac[Abbr`C`,weight_clause_FILTER,weight_clause_nub]>> - rw[]>>simp[weight_clause_def,iSUM_def] + rw[]>>simp[weight_clause_def,iSUM_def,sat_clause_alt] >- ( Cases_on`C`>>fs[]>> IF_CASES_TAC>> @@ -304,7 +318,7 @@ QED Theorem encode_correct_pbf_cnf: wfml_to_pbf wfml = (obj,pbf) ∧ satisfies w (set pbf) ⇒ - satisfies_hard (w o INL) wfml ∧ + sat_hard (w o INL) wfml ∧ &(SUM (MAP (weight_clause (w o INL)) wfml)) ≤ eval_obj obj w Proof @@ -313,7 +327,7 @@ Proof (* All hard constraints are satisfied *) gvs[wfml_to_pbf_def]>> fs[pbcTheory.satisfies_def,MEM_FLAT,MEM_MAP,PULL_EXISTS]>> - rw[satisfies_hard_def]>> + rw[sat_hard_def,sat_clause_alt]>> fs[MEM_EL]>>rw[]>>fs[LENGTH_enumerate,PULL_EXISTS]>> first_x_assum drule>> DEP_REWRITE_TAC[EL_enumerate]>> @@ -350,7 +364,7 @@ QED the sum of weights of unsatisfied clauses *) Theorem encode_correct_cnf_pbf: wfml_to_pbf wfml = (obj,pbf) ∧ - satisfies_hard w wfml ⇒ + sat_hard w wfml ⇒ ∃w'. satisfies w' (set pbf) ∧ eval_obj obj w' = &(SUM (MAP (weight_clause w) wfml)) @@ -363,7 +377,7 @@ Proof | INR y => satisfies_clause w (interp_cclause (SND (EL (y - k) wfml)))`>> CONJ_TAC >- ( - fs[satisfies_hard_def,pbcTheory.satisfies_def]>> + fs[sat_hard_def,sat_clause_alt,pbcTheory.satisfies_def]>> rw[MEM_FLAT,MEM_MAP]>> fs[MEM_EL]>>rw[]>>fs[LENGTH_enumerate,PULL_EXISTS]>> pop_assum mp_tac>> @@ -406,7 +420,7 @@ Proof unabbrev_all_tac>>CONJ_TAC >- ( Cases_on`h`>> - simp[wclause_to_pbc_def,weight_clause_def,iSUM_def]>> + simp[wclause_to_pbc_def,weight_clause_def,iSUM_def,sat_clause_alt]>> qmatch_goalsub_abbrev_tac`LENGTH C = 1`>> `wf_clause C` by fs[Abbr`C`,wf_clause_def,MEM_FILTER]>> @@ -431,7 +445,7 @@ Proof simp[EL_CONS,PRE_SUB1] QED -(* Prove injectivity of the abstract -> concrete variable map *) +(* Prove injectivity of abstract -> concrete variable map *) Theorem enc_string_INJ: INJ enc_string UNIV UNIV Proof @@ -451,12 +465,12 @@ Theorem full_encode_sem_concl: sem_concl (set pbf) obj concl ∧ conv_concl concl = SOME (SOME (lbg, ubg)) ⇒ (case lbg of - NONE => unsatisfiable_hard wfml - | SOME lb => (∀w. satisfies_hard w wfml ⇒ lb ≤ weight w wfml)) ∧ + NONE => ¬∃w. sat_hard w wfml + | SOME lb => (∀w. sat_hard w wfml ⇒ lb ≤ cost w wfml)) ∧ (case ubg of NONE => T | SOME ub => - ∃w. satisfies_hard w wfml ∧ weight w wfml ≤ ub) + ∃w. sat_hard w wfml ∧ cost w wfml ≤ ub) Proof strip_tac>> gvs[full_encode_def]>> @@ -479,31 +493,31 @@ Proof Cases_on`lbi`>>fs[] >- ( (* If the lower bound is NONE, then UNSAT *) - rw[unsatisfiable_hard_def]>> + rw[]>> CCONTR_TAC>> fs[pbcTheory.unsatisfiable_def,pbcTheory.satisfiable_def]>> metis_tac[LESS_EQ_REFL])>> rw[]>> first_x_assum drule>>rw[]>> first_x_assum drule>>rw[]>> - simp[weight_def]>>rw[nn_int_def]>> + simp[cost_def]>>rw[nn_int_def]>> intLib.ARITH_TAC)>> (* Upper bound from PB optimization *) qpat_x_assum`_ lbi _ _` kall_tac>> every_case_tac>>fs[]>> drule_all encode_correct_pbf_cnf>>rw[]>> first_x_assum (irule_at Any)>> - rw[nn_int_def,weight_def]>> + rw[nn_int_def,cost_def]>> intLib.ARITH_TAC QED Theorem FINITE_max_sat: - FINITE {weight w wfml| w | satisfies_hard w wfml} + FINITE {cost w wfml| w | sat_hard w wfml} Proof `FINITE (count (SUM (MAP FST wfml) + 1))` by fs[]>> drule_then match_mp_tac SUBSET_FINITE>> simp[IMAGE_DEF,SUBSET_DEF]>>rw[]>> - simp[weight_def] >> + simp[cost_def] >> simp[GSYM LE_LT1]>> match_mp_tac SUM_MAP_same_LE >> rw[EVERY_MEM,FORALL_PROD,weight_clause_def]>> @@ -525,17 +539,17 @@ Proof QED (* Special case *) -Theorem full_encode_sem_concl_min_unsat: +Theorem full_encode_sem_concl_opt_cost: full_encode wfml = (obj,pbf) ∧ sem_concl (set pbf) obj concl ∧ conv_concl concl = SOME (SOME (lbg, ubg)) ⇒ - (lbg = NONE ⇒ min_unsat wfml = NONE) ∧ - (lbg = ubg ⇒ min_unsat wfml = lbg) + (lbg = NONE ⇒ opt_cost wfml = NONE) ∧ + (lbg = ubg ⇒ opt_cost wfml = lbg) Proof strip_tac>> drule_all full_encode_sem_concl>> - Cases_on`lbg`>>fs[min_unsat_def]>> - rw[]>>gvs[unsatisfiable_hard_def]>> + Cases_on`lbg`>>fs[opt_cost_def]>> + rw[]>>gvs[]>> match_mp_tac MIN_SET_eq_intro>> rw[] >- @@ -553,8 +567,8 @@ Theorem full_encode_sem_output: pbc$sem_output (set pbf) obj bound (set pbf') obj' output ∧ conv_output bound output = SOME T ⇒ ∀v. - ((∃w. satisfies_hard w wfml ∧ weight w wfml ≤ v) ⇔ - (∃w'. satisfies_hard w' wfml' ∧ weight w' wfml' ≤ v)) + ((∃w. sat_hard w wfml ∧ cost w wfml ≤ v) ⇔ + (∃w'. sat_hard w' wfml' ∧ cost w' wfml' ≤ v)) Proof strip_tac>> gvs[full_encode_def]>> @@ -582,7 +596,7 @@ Proof drule_all encode_correct_pbf_cnf>> rw[]>> first_x_assum(irule_at Any)>> - fs[weight_def]>> + fs[cost_def]>> intLib.ARITH_TAC) >- ( drule_all encode_correct_cnf_pbf>>rw[]>> @@ -593,26 +607,26 @@ Proof drule_all encode_correct_pbf_cnf>> rw[]>> first_x_assum(irule_at Any)>> - fs[weight_def]>> + fs[cost_def]>> intLib.ARITH_TAC) QED (* rephrasing *) -Theorem full_encode_sem_output_min_unsat: +Theorem full_encode_sem_output_opt_cost: full_encode wfml = (obj,pbf) ∧ full_encode wfml' = (obj',pbf') ∧ pbc$sem_output (set pbf) obj bound (set pbf') obj' output ∧ conv_output bound output = SOME T ⇒ - min_unsat wfml = min_unsat wfml' + opt_cost wfml = opt_cost wfml' Proof rw[]>> drule_all full_encode_sem_output>> - rw[min_unsat_def]>> - fs[unsatisfiable_hard_def] + rw[opt_cost_def]>> + fs[] >- metis_tac[LESS_EQ_REFL] >- metis_tac[LESS_EQ_REFL]>> - `{weight w wfml | w | satisfies_hard w wfml} ≠ {} ∧ - {weight w wfml' | w | satisfies_hard w wfml'} ≠ {}` by + `{cost w wfml | w | sat_hard w wfml} ≠ {} ∧ + {cost w wfml' | w | sat_hard w wfml'} ≠ {}` by (rw[EXTENSION]>>metis_tac[])>> match_mp_tac MIN_SET_eq_intro>> simp[]>> @@ -620,18 +634,18 @@ Proof >- ( gvs[FORALL_AND_THM,PULL_EXISTS,AllCaseEqs(),EQ_IMP_THM]>> last_x_assum drule>> - disch_then(qspec_then`weight w'' wfml` mp_tac)>> + disch_then(qspec_then`cost w'' wfml` mp_tac)>> simp[]>>rw[]>> (drule_at_then Any match_mp_tac LESS_EQ_TRANS)>> drule MIN_SET_LEM >> rw[]>>gvs[PULL_EXISTS])>> drule MIN_SET_LEM>>rw[]>> gvs[PULL_EXISTS,FORALL_AND_THM,PULL_EXISTS,AllCaseEqs(),EQ_IMP_THM]>> - first_x_assum(qspecl_then[`weight w'' wfml'`,`w''`] mp_tac)>> + first_x_assum(qspecl_then[`cost w'' wfml'`,`w''`] mp_tac)>> simp[]>>rw[]>> first_assum (irule_at Any)>> - rename1`weight ww wfml`>> - first_x_assum(qspecl_then[`weight ww wfml`,`ww`] mp_tac)>> + rename1`cost ww wfml`>> + first_x_assum(qspecl_then[`cost ww wfml`,`ww`] mp_tac)>> simp[]>>rw[]>> first_x_assum drule>> simp[] From 3be550bb6e02230d7fdd4ff087654c8b4b3e3c33 Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sun, 4 Feb 2024 00:58:49 +0800 Subject: [PATCH 28/38] Revert "check SORTED before SORTING" This reverts commit 330c277822217e695cd4c01a173790eb8191c89c. --- examples/pseudo_bool/array/npbc_parseProgScript.sml | 1 - examples/pseudo_bool/pbc_normaliseScript.sml | 6 ++---- 2 files changed, 2 insertions(+), 5 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_parseProgScript.sml b/examples/pseudo_bool/array/npbc_parseProgScript.sml index a8974dfee3..04771206f1 100644 --- a/examples/pseudo_bool/array/npbc_parseProgScript.sml +++ b/examples/pseudo_bool/array/npbc_parseProgScript.sml @@ -2032,7 +2032,6 @@ val res = translate parse_assg_def; val res = translate parse_obj_term_def; val res = translate strip_obju_end_def; -val res = translate SORTED_DEF; val res = translate normalise_obj_def; val res = translate parse_obj_term_npbc_def; diff --git a/examples/pseudo_bool/pbc_normaliseScript.sml b/examples/pseudo_bool/pbc_normaliseScript.sml index 13a48db089..3f4d5f72c2 100644 --- a/examples/pseudo_bool/pbc_normaliseScript.sml +++ b/examples/pseudo_bool/pbc_normaliseScript.sml @@ -1125,9 +1125,7 @@ QED Definition normalise_obj_def: (normalise_obj NONE = NONE) ∧ (normalise_obj (SOME (f,c)) = - let f = - (if SORTED term_le f then f else QSORT term_le f) in - let (f',c') = compact_lhs f 0 in + let (f',c') = compact_lhs (QSORT term_le f) 0 in let (f'', c'') = normalise_lhs f' [] 0 in SOME (f'',c + c'+c'')) End @@ -1152,7 +1150,7 @@ Proof disch_then(qspec_then`w` assume_tac)>> drule compact_lhs_sound>> disch_then(qspec_then`w` assume_tac)>> - every_case_tac>>fs[]>> + fs[]>> intLib.ARITH_TAC QED From 99adbc667a51ac19767c30ff4eaad9200792247d Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sun, 4 Feb 2024 00:59:29 +0800 Subject: [PATCH 29/38] Revert "make translation work" This reverts commit a9a9983b1110ad892471da6913c45fae1ce43f73. --- .../array/npbc_arrayProgScript.sml | 58 +++++++++---------- .../pseudo_bool/array/npbc_listScript.sml | 12 ++-- 2 files changed, 32 insertions(+), 38 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_arrayProgScript.sml b/examples/pseudo_bool/array/npbc_arrayProgScript.sml index 5624a0e12c..a4d551b2bc 100644 --- a/examples/pseudo_bool/array/npbc_arrayProgScript.sml +++ b/examples/pseudo_bool/array/npbc_arrayProgScript.sml @@ -1554,14 +1554,7 @@ End val res = translate list_list_insert_def; val res = translate npbcTheory.dom_subst_def; - -val r = translate byte_strsub_def; -val r = translate fast_obj_constraint_def; - -val fast_obj_constraint_side = Q.prove(` - fast_obj_constraint_side x y z ⇔ T`, - EVAL_TAC>>rw[]) |> update_precondition; - +val res = translate fast_obj_constraint_def; val res = translate fast_red_subgoals_def; val res = translate do_rso_def; @@ -3241,84 +3234,84 @@ val res = translate err_obj_check_string_def; val res = translate npbc_checkTheory.eq_obj_def; val res = translate npbc_checkTheory.check_eq_obj_def; -val res = translate upd_vomap_def; - Definition w8z_def: w8z = (0w: word8) End +Definition w8o_def: + w8o = (1w: word8) +End + val w8z_v_thm = translate w8z_def; +val w8o_v_thm = translate w8o_def; val fold_update_resize_bitset = process_topdecs` fun fold_update_resize_bitset ls acc = case ls of [] => acc - | (cn::xs) => case cn of (c,x) => + | (x::xs) => if x < Word8Array.length acc then - (Word8Array.update acc x (upd_vomap c); + (Word8Array.update acc x w8o; fold_update_resize_bitset xs acc) else let val arr = Word8Array.array (2*x+1) w8z val u = Word8Array.copy acc 0 (Word8Array.length acc) arr 0 in - (Word8Array.update arr x (upd_vomap c); + (Word8Array.update arr x w8o; fold_update_resize_bitset xs arr) end ` |> append_prog; Theorem fold_update_resize_bitset_spec: ∀ls lsv accv accls. - LIST_TYPE (PAIR_TYPE INT NUM) ls lsv + LIST_TYPE NUM ls lsv ⇒ app (p : 'ffi ffi_proj) ^(fetch_v "fold_update_resize_bitset" (get_ml_prog_state())) [lsv; accv] (W8ARRAY accv accls) (POSTv v. - W8ARRAY v (FOLDL (λacc (c,n). update_resize acc 0w (upd_vomap c) n) accls ls)) + W8ARRAY v (FOLDL (λacc i. update_resize acc w8z w8o i) accls ls)) Proof Induct>> xcf "fold_update_resize_bitset" (get_ml_prog_state ())>> gvs[LIST_TYPE_def]>>xmatch - >- (xvar>>xsimpl)>> - pairarg_tac>>fs[PAIR_TYPE_def]>> - xmatch>> + >- ( + xvar>>xsimpl)>> + assume_tac w8o_v_thm>> + assume_tac w8z_v_thm>> rpt xlet_autop>> xif >- ( - xlet_autop>> xlet_autop>> xapp>>xsimpl>> simp[update_resize_def])>> - assume_tac w8z_v_thm>> rpt xlet_autop>> xapp>>xsimpl>> - simp[update_resize_def,w8z_def] + simp[update_resize_def] QED val mk_vomap_arr = process_topdecs` fun mk_vomap_arr n fc = let val acc = Word8Array.array n w8z - val f = fst fc + val f = map_snd (fst fc) val acc = fold_update_resize_bitset f acc in Word8Array.substring acc 0 (Word8Array.length acc) end` |> append_prog; Theorem map_foldl_rel: ∀ls accA accB. - MAP (CHR ∘ w2n) accA = accB ∧ - CHR (w2n wz) = cz ∧ - (!i. CHR (w2n (wo i)) = co i) ⇒ + MAP (CHR o w2n) accA = accB ⇒ MAP (CHR ∘ w2n) - (FOLDL (λacc (c,n). update_resize acc wz (wo c) n) accA ls) = - FOLDL (λacc (c,n). update_resize acc cz (co c) n) accB ls + (FOLDL (λacc i. update_resize acc w8z w8o i) accA ls) = + FOLDL (λacc i. update_resize acc #"\^@" #"\^A" i) accB ls Proof Induct>>rw[]>> - gvs[]>> - pairarg_tac>>rw[]>>gvs[]>> - rw[update_resize_def,LUPDATE_MAP] + first_x_assum match_mp_tac>> + rw[update_resize_def,LUPDATE_MAP]>> + EVAL_TAC QED Theorem mk_vomap_arr_spec: @@ -3345,8 +3338,9 @@ Proof qmatch_goalsub_abbrev_tac`strlit B`>> qsuff_tac`A = B`>- metis_tac[]>> unabbrev_all_tac>> - ho_match_mp_tac map_foldl_rel>> - simp[]>> + simp[map_snd_def]>> + match_mp_tac map_foldl_rel>> + simp[map_replicate]>> EVAL_TAC QED diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 202298e1ee..33b0e49b71 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -2323,18 +2323,18 @@ Definition check_change_obj_list_def: End Definition upd_vomap_def: - upd_vomap c = - if c = 0 ∨ Num (ABS c) ≥ 128 then - (128w : word8) + upd_vomap c n = + if c = 0 ∨ ABS c ≥ 128 then + CHR 128 else - i2w c + let (w:word8) = i2w c in + CHR (w2n w) End Definition mk_vomap_def: mk_vomap n (f,c) = strlit (FOLDL (λacc (c,n). - update_resize acc (CHR 0) - (CHR (w2n (upd_vomap c))) n ) + update_resize acc (CHR 0) (upd_vomap c n) n ) (REPLICATE n (CHR 0)) f) End From 4a0d0f54fbe89f86c72d8a48e9d275c2aab35331 Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sun, 4 Feb 2024 00:59:40 +0800 Subject: [PATCH 30/38] Revert "make npbc_list build" This reverts commit c210cf6122412ff04981191d65cc61d4ea0b6e5e. --- .../pseudo_bool/array/npbc_listScript.sml | 279 ++++++------------ 1 file changed, 87 insertions(+), 192 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 33b0e49b71..9b2c766234 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -774,28 +774,8 @@ Definition set_indices_def: (inds, sptree$insert n rinds vimap) End -Definition byte_strsub_def: - byte_strsub ls x = - (n2w (ORD (strsub ls x))):word8 -End - -Definition vomap_rel_def: - vomap_rel obj ls ⇔ - case obj of - NONE => T - | SOME (l,n) => - ∀x. - (x < strlen ls ⇒ - let w = byte_strsub ls x in - if w = 0w then ¬MEM x (MAP SND l) - else - w ≠ 128w ⇒ - ∃i. - i < LENGTH l ∧ - FST (EL i l) = w2i w ∧ - SND (EL i l) = x ) ∧ - (~(x < strlen ls) ⇒ ¬MEM x (MAP SND l)) -End +val ow = rconc (EVAL``CHR 1``); +val zw = rconc (EVAL``CHR 0``); (* Fast substitution for obj_constraint if it is in vomap *) Definition fast_obj_constraint_def: @@ -804,140 +784,16 @@ Definition fast_obj_constraint_def: INR v => if length v = 0 then ([],0) else obj_constraint (subst_fun s) l - | INL (n,v) => + | INL (n,_) => if n < strlen vomap then - let w = byte_strsub vomap n in - if w = 0w then ([],0) - else if w = 128w then obj_constraint (subst_fun s) l + if strsub vomap n = ^zw then + ([],0) else - case v of INL b => - let c = w2i w in - ([c, n], if 0 ≤ c ⇔ b then Num (ABS c) else 0) - | _ => obj_constraint (subst_fun s) l + obj_constraint (subst_fun s) l else ([],0) End -Theorem subst_aux_id: - ∀l. - EVERY (\v. f v = NONE) (MAP SND l) ⇒ - subst_aux f l = (l,[],0) -Proof - Induct>-simp[npbcTheory.subst_aux_def]>> - Cases>> - rw[npbcTheory.subst_aux_def] -QED - -Theorem subst_lhs_id: - EVERY (\v. f v = NONE) (MAP SND l) ⇒ - subst_lhs f l = (l, 0) -Proof - rw[npbcTheory.subst_lhs_def]>> - rpt(pairarg_tac>>fs[])>> - drule subst_aux_id>>strip_tac>> - gvs[EVAL``clean_up []``]>> - Cases_on`l`>> - gvs[npbcTheory.add_lists_def] -QED - -Theorem subst_lhs_subst_fun_bool: - ∀l. - subst_aux (subst_fun (INL (x,INL b))) l = - (FILTER (λ(c,y). x ≠ y) l, [], - SUM (MAP (Num o ABS o FST) (FILTER (λ(c,y). x = y ∧ (0 ≤ c ⇔ b)) l)) - ) -Proof - Induct>>rw[npbcTheory.subst_aux_def]>> - pairarg_tac>>gvs[npbcTheory.subst_aux_def]>> - simp[subst_fun_def] -QED - -Theorem add_lists_emp_2: - add_lists ls [] = (ls,0) -Proof - Cases_on`ls`>>EVAL_TAC -QED - -Theorem add_lists_map_negate_coeff: - ∀ls rs. - rs = (MAP (λ(c,l). (-c,l)) ls) ⇒ - add_lists ls rs = ([],SUM (MAP (λi. Num (ABS (FST i))) ls)) -Proof - ho_match_mp_tac npbcTheory.add_lists_ind>> - simp[npbcTheory.add_lists_def,npbcTheory.add_terms_def] -QED - -Theorem add_lists_FILTER_map_negate_coeff: - ∀ls rs. - SORTED $< (MAP SND ls) ∧ - i < LENGTH ls ∧ - FILTER (λ(c,n). SND (EL i ls) ≠ n) (MAP (λ(c,l). (-c,l)) ls) = rs ⇒ - add_lists ls rs = - ([EL i ls], - SUM (MAP (λi. Num (ABS (FST i))) ls) - - Num (ABS (FST (EL i ls)))) -Proof - cheat -QED - -Theorem vomap_rel_fast_obj_constraint: - SORTED $< (MAP SND (FST l)) ∧ - vomap_rel (SOME l) vomap ⇒ - fast_obj_constraint s l vomap = - obj_constraint (subst_fun s) l -Proof - rw[fast_obj_constraint_def]>> - Cases_on`l`>> - fs[npbcTheory.obj_constraint_def,subst_fun_def]>> - reverse TOP_CASE_TAC>>gvs[] - >- ( - (* INR *) - rw[]>> - rpt (pairarg_tac>>fs[])>> - pop_assum mp_tac>> - DEP_REWRITE_TAC[add_lists_map_negate_coeff]>>rw[]>> - pop_assum mp_tac>> - DEP_REWRITE_TAC[subst_lhs_id]>> - simp[EVERY_MAP,subst_fun_def,EVERY_MEM])>> - TOP_CASE_TAC>>simp[]>> - fs[vomap_rel_def,FORALL_PROD,FORALL_AND_THM]>> - reverse TOP_CASE_TAC>>simp[] - >- ( - (* not in map *) - first_x_assum drule>> - rw[]>> - rpt (pairarg_tac>>fs[])>> - pop_assum mp_tac>> - DEP_REWRITE_TAC[add_lists_map_negate_coeff]>>rw[]>> - pop_assum mp_tac>> - DEP_REWRITE_TAC[subst_lhs_id]>> - simp[EVERY_MAP,subst_fun_def,EVERY_MEM]>> - simp[FORALL_PROD]>> - metis_tac[PAIR,SND,MEM_MAP])>> - first_x_assum drule>>rw[]>>fs[] - >- ( - (* not in map *) - rpt (pairarg_tac>>fs[])>> - pop_assum mp_tac>> - DEP_REWRITE_TAC[add_lists_map_negate_coeff]>>rw[]>> - pop_assum mp_tac>> - DEP_REWRITE_TAC[subst_lhs_id]>> - simp[EVERY_MAP,subst_fun_def,EVERY_MEM,FORALL_PROD]>> - rw[]>>metis_tac[PAIR,SND,MEM_MAP])>> - TOP_CASE_TAC>>simp[]>> - simp[npbcTheory.subst_lhs_def]>> - simp[subst_lhs_subst_fun_bool,npbcTheory.clean_up_def,add_lists_emp_2]>> - drule add_lists_FILTER_map_negate_coeff>> - disch_then drule>> - disch_then (fn th => DEP_REWRITE_TAC[GEN_ALL th])>> - simp[]>> - CONJ_TAC >- - metis_tac[PAIR]>> - qmatch_goalsub_abbrev_tac`A - ( _ - B + C)`>> - qmatch_goalsub_abbrev_tac`w2i bb`>> - cheat -QED - Definition fast_red_subgoals_def: fast_red_subgoals ord s def obj vomap = let cobj = @@ -1707,20 +1563,75 @@ Proof simp[any_el_ALT] QED -Definition sorted_obj_def: - sorted_obj obj ⇔ - case obj of NONE => T - | SOME (l:(int # num) list # int) => SORTED $< (MAP SND (FST l)) +Definition vomap_rel_def: + vomap_rel obj ls ⇔ + case obj of + NONE => T + | SOME l => + ∀x. + MEM x (MAP SND (FST l)) <=> + x < strlen ls ∧ strsub ls x ≠ ^zw End +Theorem add_lists_map_negate_coeff: + ∀ls rs. + rs = (MAP (λ(c,l). (-c,l)) ls) ⇒ + add_lists ls rs = ([],SUM (MAP (λi. Num (ABS (FST i))) ls)) +Proof + ho_match_mp_tac npbcTheory.add_lists_ind>> + simp[npbcTheory.add_lists_def,npbcTheory.add_terms_def] +QED + +Theorem subst_aux_id: + ∀l. + EVERY (\v. f v = NONE) (MAP SND l) ⇒ + subst_aux f l = (l,[],0) +Proof + Induct>-simp[npbcTheory.subst_aux_def]>> + Cases>> + rw[npbcTheory.subst_aux_def] +QED + +Theorem subst_lhs_id: + EVERY (\v. f v = NONE) (MAP SND l) ⇒ + subst_lhs f l = (l, 0) +Proof + rw[npbcTheory.subst_lhs_def]>> + rpt(pairarg_tac>>fs[])>> + drule subst_aux_id>>strip_tac>> + gvs[EVAL``clean_up []``]>> + Cases_on`l`>> + gvs[npbcTheory.add_lists_def] +QED + +Theorem vomap_rel_fast_obj_constraint: + vomap_rel (SOME l) vomap ⇒ + fast_obj_constraint s l vomap = + obj_constraint (subst_fun s) l +Proof + rw[fast_obj_constraint_def]>> + every_case_tac>> + Cases_on`l`>> + fs[npbcTheory.obj_constraint_def,subst_fun_def]>> + rpt (pairarg_tac>>fs[])>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[add_lists_map_negate_coeff]>>rw[]>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[subst_lhs_id]>> + fs[vomap_rel_def]>> + simp[EVERY_MAP,LAMBDA_PROD,subst_fun_def]>> + gvs[EVERY_MEM]>> + rw[]>>pairarg_tac>>fs[EXTENSION,MEM_MAP,domain_lookup]>> + metis_tac[option_CLAUSES,SND,PAIR] +QED + Theorem vomap_rel_fast_red_subgoals: - sorted_obj obj ∧ vomap_rel obj vomap ⇒ fast_red_subgoals ord s def obj vomap = red_subgoals ord (subst_fun s) def obj Proof rw[fast_red_subgoals_def,red_subgoals_def]>> - every_case_tac>>fs[sorted_obj_def]>> + every_case_tac>>fs[]>> metis_tac[vomap_rel_fast_obj_constraint] QED @@ -1728,7 +1639,6 @@ Theorem fml_rel_check_red_list: fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ vimap_rel fmlls vimap ∧ - sorted_obj obj ∧ vomap_rel obj vomap ∧ (∀n. n ≥ id ⇒ any_el n fmlls NONE = NONE) ∧ check_red_list ord obj b tcb fmlls inds id c s pfs @@ -2154,7 +2064,6 @@ Theorem fml_rel_check_sstep_list: fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ vimap_rel fmlls vimap ∧ - sorted_obj obj ∧ vomap_rel obj vomap ∧ (∀n. n ≥ id ⇒ any_el n fmlls NONE = NONE) ∧ check_sstep_list sstep ord obj tcb fmlls inds id vimap vomap = @@ -2322,32 +2231,33 @@ Definition check_change_obj_list_def: else NONE) End -Definition upd_vomap_def: - upd_vomap c n = - if c = 0 ∨ ABS c ≥ 128 then - CHR 128 - else - let (w:word8) = i2w c in - CHR (w2n w) -End - Definition mk_vomap_def: mk_vomap n (f,c) = - strlit (FOLDL (λacc (c,n). - update_resize acc (CHR 0) (upd_vomap c n) n ) - (REPLICATE n (CHR 0)) f) + strlit (FOLDL (λacc i. update_resize acc ^zw ^ow i) (REPLICATE n ^zw) (MAP SND f)) End +Theorem resize_acc_bitset_iff: + ∀ls acc. + (x < + LENGTH + (FOLDL (λacc i. update_resize acc ^zw ^ow i) acc ls) ∧ + EL x (FOLDL (λacc i. update_resize acc ^zw ^ow i) acc ls) ≠ ^zw) ⇔ + (MEM x ls ∨ x < LENGTH acc ∧ EL x acc ≠ ^zw) +Proof + Induct>>rw[]>> + rw[update_resize_def,EL_LUPDATE,EL_APPEND_EQN,EL_REPLICATE]>> + EVERY_CASE_TAC>>gvs[]>> + Cases_on`x < 2 * h + 1` >>simp[]>> + DEP_REWRITE_TAC[EL_REPLICATE]>> + simp[] +QED + Theorem vomap_rel_mk_vomap: - SORTED $< (MAP SND (FST fc)) ⇒ vomap_rel (SOME fc) (mk_vomap n fc) Proof Cases_on`fc`>>rw[vomap_rel_def,mk_vomap_def]>> - rw[]>> - cheat - (* simp[resize_acc_bitset_iff]>> - metis_tac[EL_REPLICATE] *) + metis_tac[EL_REPLICATE] QED Definition check_cstep_list_def: @@ -2626,7 +2536,6 @@ Theorem fml_rel_check_cstep_list: fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ vimap_rel fmlls vimap ∧ - sorted_obj pc.obj ∧ vomap_rel pc.obj vomap ∧ (∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE) ∧ check_cstep_list cstep fmlls inds vimap vomap pc = @@ -2636,7 +2545,6 @@ Theorem fml_rel_check_cstep_list: fml_rel fml' fmlls' ∧ ind_rel fmlls' inds' ∧ vimap_rel fmlls' vimap' ∧ - sorted_obj pc'.obj ∧ vomap_rel pc'.obj vomap' ∧ (∀n. n ≥ pc'.id ⇒ any_el n fmlls' NONE = NONE) ∧ pc.id ≤ pc'.id @@ -2843,11 +2751,8 @@ Proof metis_tac[ind_rel_reindex])>> CONJ_TAC >- metis_tac[fml_rel_fml_rel_vimap_rel]>> - CONJ_ASM1_TAC >- - cheat>> - CONJ_TAC >- ( - fs[sorted_obj_def]>> - metis_tac[vomap_rel_mk_vomap])>> + CONJ_TAC >- + metis_tac[vomap_rel_mk_vomap]>> simp[any_el_rollback]) >- ( (* CheckObj *) fs[check_cstep_def,check_cstep_list_def] @@ -2870,7 +2775,6 @@ Theorem fml_rel_check_csteps_list: fml_rel fml fmlls ∧ ind_rel fmlls inds ∧ vimap_rel fmlls vimap ∧ - sorted_obj pc.obj ∧ vomap_rel pc.obj vomap ∧ (∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE) ∧ check_csteps_list csteps fmlls inds vimap vomap pc = @@ -2880,7 +2784,6 @@ Theorem fml_rel_check_csteps_list: fml_rel fml' fmlls' ∧ ind_rel fmlls' inds' ∧ vimap_rel fmlls' vimap' ∧ - sorted_obj pc'.obj ∧ vomap_rel pc'.obj vomap' ∧ (∀n. n ≥ pc'.id ⇒ any_el n fmlls' NONE = NONE) ∧ pc.id ≤ pc'.id @@ -3118,7 +3021,6 @@ Definition mk_vomap_opt_def: End Theorem check_csteps_list_concl: - sorted_obj obj ∧ check_csteps_list cs (FOLDL (λacc (i,v). update_resize acc NONE (SOME (v,T)) i) (REPLICATE m NONE) (enumerate 1 fml)) @@ -3144,14 +3046,11 @@ Proof `vimap_rel fmlls vimap` by ( unabbrev_all_tac>> simp[vimap_rel_FOLDL_update_resize])>> - `sorted_obj pc.obj` by - gvs[sorted_obj_def,init_conf_def,Abbr`pc`]>> `vomap_rel pc.obj vomap` by ( unabbrev_all_tac>> simp[init_conf_def]>> Cases_on`obj`>- EVAL_TAC>> - gvs[sorted_obj_def]>> metis_tac[vomap_rel_mk_vomap,mk_vomap_opt_def])>> `∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE` by ( rw[Abbr`pc`,Abbr`fmlls`,any_el_ALT,init_conf_def]>> @@ -3244,7 +3143,6 @@ Proof QED Theorem check_csteps_list_output: - sorted_obj obj ∧ check_csteps_list cs (FOLDL (λacc (i,v). update_resize acc NONE (SOME (v,T)) i) (REPLICATE m NONE) (enumerate 1 fml)) @@ -3269,14 +3167,11 @@ Proof `vimap_rel fmlls vimap` by ( unabbrev_all_tac>> simp[vimap_rel_FOLDL_update_resize])>> - `sorted_obj pc.obj` by - gvs[sorted_obj_def,init_conf_def,Abbr`pc`]>> `vomap_rel pc.obj vomap` by ( unabbrev_all_tac>> simp[init_conf_def]>> Cases_on`obj`>- EVAL_TAC>> - gvs[sorted_obj_def]>> metis_tac[vomap_rel_mk_vomap,mk_vomap_opt_def])>> `∀n. n ≥ pc.id ⇒ any_el n fmlls NONE = NONE` by ( rw[Abbr`pc`,Abbr`fmlls`,any_el_ALT,init_conf_def]>> From b84bf2b58631ec6c655660475c814b47efed5591 Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sun, 4 Feb 2024 02:01:45 +0800 Subject: [PATCH 31/38] faster change obj check --- .../array/npbc_arrayProgScript.sml | 18 ++++++------- .../pseudo_bool/array/npbc_listScript.sml | 6 ++--- examples/pseudo_bool/npbc_checkScript.sml | 27 ++++++++++++++++--- 3 files changed, 35 insertions(+), 16 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_arrayProgScript.sml b/examples/pseudo_bool/array/npbc_arrayProgScript.sml index a4d551b2bc..2caa1a3a99 100644 --- a/examples/pseudo_bool/array/npbc_arrayProgScript.sml +++ b/examples/pseudo_bool/array/npbc_arrayProgScript.sml @@ -2970,6 +2970,7 @@ val res = translate emp_vec_def; val res = translate do_change_obj_check_def; val res = translate npbc_checkTheory.add_obj_def; val res = translate npbc_checkTheory.mk_diff_obj_def; +val res = translate npbc_checkTheory.mk_tar_obj_def; val check_change_obj_arr = process_topdecs ` fun check_change_obj_arr lno b fml id obj fc' pfs = @@ -2977,16 +2978,15 @@ val check_change_obj_arr = process_topdecs ` raise Fail (format_failure lno ("no objective to change")) | Some fc => let - val fc' = mk_diff_obj b fc fc' - val csubs = change_obj_subgoals fc fc' - val b = True + val csubs = change_obj_subgoals (mk_tar_obj b fc) fc' + val bb = True val e = [] - val cpfs = extract_clauses_arr lno emp_vec b fml csubs pfs e in - case check_subproofs_arr lno cpfs b fml id id of + val cpfs = extract_clauses_arr lno emp_vec bb fml csubs pfs e in + case check_subproofs_arr lno cpfs bb fml id id of (fml', id') => let val u = rollback_arr fml' id id' in if do_change_obj_check pfs csubs then - (fml',(fc',id')) + (fml',(mk_diff_obj b fc fc',id')) else raise Fail (format_failure lno ("Objective change subproofs did not cover all subgoals. Expected: #[1-2]")) end end` |> append_prog @@ -3039,14 +3039,14 @@ Proof (λv. ARRAY aa vv * &(case extract_clauses_list emp_vec T fmlls - (change_obj_subgoals x (mk_diff_obj b x fc)) pfs [] of + (change_obj_subgoals (mk_tar_obj b x) fc) pfs [] of NONE => F | SOME res => LIST_TYPE (PAIR_TYPE (OPTION_TYPE (PAIR_TYPE (LIST_TYPE constraint_TYPE) NUM)) (LIST_TYPE NPBC_CHECK_LSTEP_TYPE)) res v)) (λe. ARRAY aa vv * & (Fail_exn e ∧ extract_clauses_list emp_vec T fmlls - (change_obj_subgoals x (mk_diff_obj b x fc)) pfs [] = NONE)))` + (change_obj_subgoals (mk_tar_obj b x) fc) pfs [] = NONE)))` >- ( xapp>>xsimpl>> CONJ_TAC >- EVAL_TAC>> @@ -3090,7 +3090,7 @@ Proof rpt xlet_autop>> xif >- ( - xlet_autop>> + rpt xlet_autop>> xcon>>xsimpl>> fs[PAIR_TYPE_def]>> metis_tac[ARRAY_refl])>> diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 9b2c766234..578550de24 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -2217,8 +2217,7 @@ Definition check_change_obj_list_def: check_change_obj_list b fml id obj fc' pfs ⇔ case obj of NONE => NONE | SOME fc => - let fc' = mk_diff_obj b fc fc' in - let csubs = change_obj_subgoals fc fc' in + let csubs = change_obj_subgoals (mk_tar_obj b fc) fc' in case extract_clauses_list emp_vec T fml csubs pfs [] of NONE => NONE | SOME cpfs => @@ -2227,7 +2226,8 @@ Definition check_change_obj_list_def: | SOME (fml',id') => let rfml = rollback fml' id id' in if do_change_obj_check pfs csubs then - SOME (rfml,fc',id') + let fc'' = mk_diff_obj b fc fc' in + SOME (rfml,fc'',id') else NONE) End diff --git a/examples/pseudo_bool/npbc_checkScript.sml b/examples/pseudo_bool/npbc_checkScript.sml index 4c1101febc..9ea16a9333 100644 --- a/examples/pseudo_bool/npbc_checkScript.sml +++ b/examples/pseudo_bool/npbc_checkScript.sml @@ -2024,12 +2024,17 @@ Definition mk_diff_obj_def: if b then fc' else add_obj fc fc' End +Definition mk_tar_obj_def: + mk_tar_obj b fc = + (if b then fc else ([],0:int)) +End + Definition check_change_obj_def: check_change_obj b fml id obj fc' pfs ⇔ case obj of NONE => NONE | SOME fc => - ( let fc' = mk_diff_obj b fc fc' in - let csubs = change_obj_subgoals fc fc' in + ( let csubs = change_obj_subgoals + (mk_tar_obj b fc) fc' in case extract_clauses (λx. NONE) T fml csubs pfs [] of NONE => NONE | SOME cpfs => @@ -2044,7 +2049,8 @@ Definition check_change_obj_def: ) (enumerate 0 csubs) then - SOME (fc',id') + let fc'' = mk_diff_obj b fc fc' in + SOME (fc'',id') else NONE)) End @@ -3034,7 +3040,7 @@ Proof strip_tac>> fs[EVERY_MEM,MEM_MAP,EXISTS_PROD]>> gvs[change_obj_subgoals_def,MEM_enumerate_iff,ADD1,PULL_EXISTS]>> - qmatch_asmsub_rename_tac`model_bounding fnew fold`>> + qmatch_asmsub_abbrev_tac`model_bounding fnew fold`>> `unsatisfiable (core_only_fml T fml ∪ {not (model_bounding fnew fold)})` by ( first_x_assum(qspec_then`0` mp_tac)>> simp[]>> @@ -3067,6 +3073,7 @@ Proof drule check_contradiction_unsat>> simp[unsatisfiable_def,satisfiable_def] ))>> + `∀w. satisfies w (core_only_fml T fml) ⇒ eval_obj (SOME fold) w = @@ -3076,6 +3083,18 @@ Proof rw[]>>gvs[]>> ntac 2 (first_x_assum (qspec_then `w` mp_tac))>>simp[]>> intLib.ARITH_TAC)>> + + qmatch_goalsub_abbrev_tac`SOME fupd`>> + `∀w. + satisfies w (core_only_fml T fml) ⇒ + eval_obj (SOME x) w = + eval_obj (SOME fupd) w` by + (unabbrev_all_tac>>gvs[mk_diff_obj_def,mk_tar_obj_def]>> + rw[]>>first_x_assum drule>> + rw[]>> + simp[add_obj_eval_obj]>> + EVAL_TAC)>> + gvs[]>> CONJ_TAC >- fs[id_ok_def]>> CONJ_TAC >- ( From c82f0e4dca793703815ffa5ab8a6be9499eebb97 Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sun, 4 Feb 2024 14:42:38 +0800 Subject: [PATCH 32/38] use mergesort --- .../array/npbc_parseProgScript.sml | 45 +++++++++++++++++++ examples/pseudo_bool/pbc_normaliseScript.sml | 14 +++--- 2 files changed, 52 insertions(+), 7 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_parseProgScript.sml b/examples/pseudo_bool/array/npbc_parseProgScript.sml index 04771206f1..ec291a44b8 100644 --- a/examples/pseudo_bool/array/npbc_parseProgScript.sml +++ b/examples/pseudo_bool/array/npbc_parseProgScript.sml @@ -73,6 +73,51 @@ val r = translate term_le_def; val r = translate mk_coeff_def; val r = translate normalise_lhs_def; +val r = translate mergesortTheory.sort2_def; +val r = translate mergesortTheory.sort3_def; +val r = translate mergesortTheory.merge_def; +val r = translate DROP_def; +val r = translate (mergesortTheory.mergesortN_def |> SIMP_RULE std_ss [DIV2_def]); + +Triviality mergesortn_ind: + mergesortn_ind (:'a) +Proof + once_rewrite_tac [fetch "-" "mergesortn_ind_def"] + \\ rpt gen_tac + \\ rpt (disch_then strip_assume_tac) + \\ match_mp_tac (latest_ind ()) + \\ rpt strip_tac + \\ last_x_assum match_mp_tac + \\ rpt strip_tac + \\ gvs [FORALL_PROD, DIV2_def] +QED + +val _ = mergesortn_ind |> update_precondition; + +val r = translate mergesortTheory.mergesort_def; + +val r = translate mergesortTheory.sort2_tail_def; +val r = translate mergesortTheory.sort3_tail_def; +val r = translate mergesortTheory.merge_tail_def; +val r = translate (mergesortTheory.mergesortN_tail_def |> SIMP_RULE std_ss [DIV2_def]); + +Triviality mergesortn_tail_ind: + mergesortn_tail_ind (:'a) +Proof + once_rewrite_tac [fetch "-" "mergesortn_tail_ind_def"] + \\ rpt gen_tac + \\ rpt (disch_then strip_assume_tac) + \\ match_mp_tac (latest_ind ()) + \\ rpt strip_tac + \\ last_x_assum match_mp_tac + \\ rpt strip_tac + \\ gvs [FORALL_PROD, DIV2_def] +QED + +val _ = mergesortn_tail_ind |> update_precondition; + +val r = translate mergesortTheory.mergesort_tail_def; + val r = translate pbc_to_npbc_def; val pbc_to_npbc_side = Q.prove( `∀x. pbc_to_npbc_side x <=> T`, diff --git a/examples/pseudo_bool/pbc_normaliseScript.sml b/examples/pseudo_bool/pbc_normaliseScript.sml index 3f4d5f72c2..e4d2bedada 100644 --- a/examples/pseudo_bool/pbc_normaliseScript.sml +++ b/examples/pseudo_bool/pbc_normaliseScript.sml @@ -1,7 +1,7 @@ (* Normalizes pbc into npbc *) -open preamble pbcTheory npbcTheory mlmapTheory; +open preamble pbcTheory npbcTheory mlmapTheory mergesortTheory; val _ = new_theory "pbc_normalise"; @@ -407,13 +407,13 @@ Proof intLib.ARITH_TAC QED -Theorem iSUM_QSORT_term_le[simp]: - iSUM (MAP (eval_term w) (QSORT $≤ l)) = +Theorem iSUM_mergesort_term_le[simp]: + iSUM (MAP (eval_term w) (mergesort $≤ l)) = iSUM (MAP (eval_term w) l) Proof match_mp_tac iSUM_PERM>> match_mp_tac PERM_MAP>> - metis_tac[QSORT_PERM,PERM_SYM] + metis_tac[mergesort_perm,PERM_SYM] QED Theorem eval_lit_eq_flip: @@ -510,7 +510,7 @@ QED Definition pbc_to_npbc_def: (pbc_to_npbc (GreaterEqual,lhs,n) = - let (lhs',m') = compact_lhs (QSORT term_le lhs) 0 in + let (lhs',m') = compact_lhs (mergesort term_le lhs) 0 in let (lhs'',m'') = normalise_lhs lhs' [] 0 in let rhs = if n-(m'+m'') ≥ 0 then Num(n-(m'+m'')) else 0 in (lhs'',rhs):npbc) ∧ @@ -648,7 +648,7 @@ Proof imp_res_tac compact_lhs_no_dup>> pop_assum mp_tac>> impl_tac>- ( - match_mp_tac QSORT_SORTED>> + match_mp_tac mergesort_sorted>> fs[transitive_term_le]>> simp[total_def]>> Cases>>Cases>>simp[])>> @@ -1125,7 +1125,7 @@ QED Definition normalise_obj_def: (normalise_obj NONE = NONE) ∧ (normalise_obj (SOME (f,c)) = - let (f',c') = compact_lhs (QSORT term_le f) 0 in + let (f',c') = compact_lhs (mergesort term_le f) 0 in let (f'', c'') = normalise_lhs f' [] 0 in SOME (f'',c + c'+c'')) End From cf1d94467defcfd451087f4cfb844f781795c38e Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sun, 4 Feb 2024 18:18:16 +0800 Subject: [PATCH 33/38] onepass obj_constraint --- .../array/npbc_arrayProgScript.sml | 9 +- .../pseudo_bool/array/npbc_listScript.sml | 269 +++++++++++++++++- 2 files changed, 273 insertions(+), 5 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_arrayProgScript.sml b/examples/pseudo_bool/array/npbc_arrayProgScript.sml index 2caa1a3a99..5317733473 100644 --- a/examples/pseudo_bool/array/npbc_arrayProgScript.sml +++ b/examples/pseudo_bool/array/npbc_arrayProgScript.sml @@ -57,8 +57,8 @@ val res = translate foldi_def; val res = translate toAList_def; val r = translate add_terms_def; -val r = translate add_lists'_def; -val r = translate add_lists'_thm; +val r = translate add_listsLR_def; +val r = translate add_listsLR_thm; val r = translate add_def; val r = translate multiply_def; @@ -1554,6 +1554,11 @@ End val res = translate list_list_insert_def; val res = translate npbcTheory.dom_subst_def; + +val res = translate obj_single_aux_def; +val res = translate obj_single_def; +val res = translate full_obj_single_def; + val res = translate fast_obj_constraint_def; val res = translate fast_red_subgoals_def; val res = translate do_rso_def; diff --git a/examples/pseudo_bool/array/npbc_listScript.sml b/examples/pseudo_bool/array/npbc_listScript.sml index 578550de24..f5149e30f2 100644 --- a/examples/pseudo_bool/array/npbc_listScript.sml +++ b/examples/pseudo_bool/array/npbc_listScript.sml @@ -774,23 +774,286 @@ Definition set_indices_def: (inds, sptree$insert n rinds vimap) End +Definition add_listsLR_def: + (add_listsL cx (xn:num) xs1 ys zs n = + case ys of + | [] => (REV zs ((cx,xn)::xs1),n) + | (y::ys1) => + let (cy,yn) = y in + if yn < xn then add_listsL cx xn xs1 ys1 (y::zs) n + else + if xn < yn then add_listsR cy yn xs1 ys1 ((cx,xn)::zs) n + else + let (zs1,n1) = add_terms cx cy xn zs n in + add_listsLR xs1 ys1 zs1 n1) ∧ + (add_listsR cy yn xs ys1 zs n = + case xs of + | [] => (REV zs ((cy,yn)::ys1),n) + | (x::xs1) => + let (cx,xn) = x in + if xn < yn then add_listsR cy yn xs1 ys1 (x::zs) n + else + if yn < xn then add_listsL cx xn xs1 ys1 ((cy,yn)::zs) n + else + let (zs1,n1) = add_terms cx cy xn zs n in + add_listsLR xs1 ys1 zs1 n1) ∧ + (add_listsLR xs ys zs n = + case xs of + | [] => (REV zs ys,n) + | (x::xs1) => + case ys of + | [] => (REV zs xs,n) + | (y::ys1) => + let (cx,xn) = x in + let (cy,yn) = y in + if xn < yn then add_listsR cy yn xs1 ys1 (x::zs) n + else + if yn < xn then add_listsL cx xn xs1 ys1 (y::zs) n + else + let (zs1,n1) = add_terms cx cy xn zs n in + add_listsLR xs1 ys1 zs1 n1) +Termination + WF_REL_TAC `measure (λv. + case v of + INL (cx,xn,xs1,ys,zs,n) => LENGTH xs1 + LENGTH ys + | INR v => + case v of + INL (cy,yn,xs,ys1,zs,n) => LENGTH xs + LENGTH ys1 + | INR (xs,ys,zs,n) => LENGTH xs + LENGTH ys)`>> + rw[] +End + +Theorem add_listsLR_eq: + (∀cx xn xs1 ys zs n. + add_listsL cx xn xs1 ys zs n = + add_lists' ((cx,xn)::xs1) ys zs n) ∧ + (∀cy yn xs ys1 zs n. + add_listsR cy yn xs ys1 zs n = + add_lists' xs ((cy,yn)::ys1) zs n) ∧ + (∀xs ys zs n. + add_lists' xs ys zs n = + add_listsLR xs ys zs n) +Proof + ho_match_mp_tac add_listsLR_ind>>rw[]>> + simp[Once npbcTheory.add_lists'_def,Once add_listsLR_def]>> + every_case_tac>>simp[]>> + rpt(pairarg_tac>>gvs[])>>rw[]>>fs[] +QED + +Theorem add_listsLR_thm: + add_lists xs ys = add_listsLR xs ys [] 0 +Proof + rw[npbcTheory.add_lists'_thm,add_listsLR_eq] +QED + val ow = rconc (EVAL``CHR 1``); val zw = rconc (EVAL``CHR 0``); +Theorem subst_aux_no_INR_FILTER: + ∀l. + EVERY (λ(c,x). case f x of SOME (INR _ ) => F | _ => T) l ⇒ + subst_aux f l = + (FILTER (λ(c,x). f x = NONE) l, + [], + SUM (MAP (λ(c,x). + if is_Pos c ⇔ OUTL (THE (f x)) + then Num(ABS c) else 0) + (FILTER (λ(c,x). f x ≠ NONE) l))) +Proof + Induct>>rw[npbcTheory.subst_aux_def]>> + rpt(pairarg_tac>>gvs[])>> + simp[npbcTheory.subst_aux_def]>> + Cases_on`f x`>>fs[]>> + Cases_on`x'`>>fs[]>> + rw[]>>fs[] +QED + +Theorem add_lists_emp_2: + add_lists ls [] = (ls,0) +Proof + Cases_on`ls`>>EVAL_TAC +QED + +Theorem subst_lhs_no_INR_FILTER: + ∀l. + EVERY (λ(c,x). case f x of SOME (INR _ ) => F | _ => T) l ⇒ + subst_lhs f l = + (FILTER (λ(c,x). f x = NONE) l, + SUM (MAP (λ(c,x). + if is_Pos c ⇔ OUTL (THE (f x)) + then Num(ABS c) else 0) + (FILTER (λ(c,x). f x ≠ NONE) l))) +Proof + rw[npbcTheory.subst_lhs_def]>> + drule subst_aux_no_INR_FILTER>> + rw[]>> + simp[npbcTheory.clean_up_def,add_lists_emp_2] +QED + +Theorem SORTED_add_lists_FILTER_MAP: + ∀l. + SORTED $< (MAP SND l) ⇒ + add_lists l + (FILTER (λ(c,x). f x = NONE) (MAP (λ(c,l). (-c,l)) l)) = + (FILTER (λ(c,x). f x <> NONE) l, + SUM (MAP (λ(c,x). Num (ABS c)) + (FILTER (λ(c,x). f x = NONE) l)) ) +Proof + Induct>>rw[npbcTheory.add_lists_def]>> + rpt(pairarg_tac>>gvs[])>> + drule SORTED_TL>>rw[]>>gvs[]>> + simp[npbcTheory.add_lists_def] + >- (EVAL_TAC>>rw[])>> + qmatch_goalsub_abbrev_tac`add_lists _ lss`>> + Cases_on`lss`>>gvs[add_lists_emp_2]>> + Cases_on`h`>>gvs[npbcTheory.add_lists_def]>> + `MEM r (MAP SND l)` by + (pop_assum (mp_tac o Q.AP_TERM `λx. MEM r (MAP SND x)`)>> + simp[MEM_MAP,MEM_FILTER,PULL_EXISTS,EXISTS_PROD]>> + metis_tac[])>> + `x < r` by ( + qpat_x_assum`SORTED _ (_ :: _)` mp_tac>> + DEP_REWRITE_TAC[SORTED_EQ]>> + simp[])>> + simp[] +QED + +Theorem SUM_SPLIT_LE: + ∀l. + SUM (MAP (λ(c,x). Num (ABS c)) (FILTER (λ(c,x). f x = NONE) l)) + + SUM + (MAP (λ(c,x). if 0 ≤ c ⇔ OUTL (THE (f x)) then Num (ABS c) else 0) + (FILTER (λ(c,x). f x ≠ NONE) (MAP (λ(c,l). (-c,l)) l))) ≤ + SUM (MAP (λi. Num (ABS (FST i))) l) +Proof + Induct>>simp[FORALL_PROD]>>rw[]>> + rw[] +QED + +Theorem obj_constraint_simp: + EVERY (λ(c,x). case f x of SOME (INR _ ) => F | _ => T) l ∧ + SORTED $< (MAP SND l) ⇒ + obj_constraint f (l,b) = + (FILTER (λ(c,x). f x ≠ NONE) l, + SUM + (MAP + (λ(c,x). if 0 ≤ c ⇔ OUTL (THE (f x)) then Num (ABS c) else 0) + (FILTER (λ(c,x). f x ≠ NONE) l))) +Proof + rw[npbcTheory.obj_constraint_def]>> + DEP_REWRITE_TAC[subst_lhs_no_INR_FILTER]>> + CONJ_TAC >- ( + gvs[EVERY_MEM,MEM_MAP,PULL_EXISTS,FORALL_PROD]>> + metis_tac[])>> + simp[]>> + DEP_REWRITE_TAC[SORTED_add_lists_FILTER_MAP]>> + simp[]>> + pop_assum kall_tac>> + pop_assum mp_tac>> + Induct_on`l`>>simp[FORALL_PROD]>> + rw[]>> + every_case_tac>>gvs[] + >- ( + `p_1 = 0` by intLib.ARITH_TAC>> + simp[]) + >- ( + DEP_REWRITE_TAC[LESS_EQ_ADD_SUB]>> + simp[SUM_SPLIT_LE]) + >- ( + DEP_REWRITE_TAC[LESS_EQ_ADD_SUB]>> + simp[SUM_SPLIT_LE]>> + Cases_on`x'`>>fs[]>> + intLib.ARITH_TAC) +QED + +(* one pass obj_constraint *) +Definition obj_single_aux_def: + (obj_single_aux f n [] acc k = SOME(REVERSE acc,k:num)) ∧ + (obj_single_aux f n ((c,x:num)::xs) acc k = + if n < x then + case f x of + NONE => obj_single_aux f x xs acc k + | SOME (INL b) => + let r = if is_Pos c ⇔ b then k + Num (ABS c) else k in + obj_single_aux f x xs ((c,x)::acc) r + | SOME (INR _) => NONE + else NONE) +End + +Definition obj_single_def: + (obj_single f [] = SOME([],0:num)) ∧ + (obj_single f ((c,x:num)::xs) = + case f x of + NONE => obj_single_aux f x xs [] 0 + | SOME (INL b) => + let r = if is_Pos c ⇔ b then Num (ABS c) else 0 in + obj_single_aux f x xs [(c,x)] r + | SOME (INR _) => NONE) +End + +Theorem obj_single_aux_eq_SOME: + ∀l f n acc k res. + obj_single_aux f n l acc k = SOME res ⇒ + EVERY (λ(c,x). case f x of SOME (INR _ ) => F | _ => T) l ∧ + SORTED $< (n::MAP SND l) ∧ + res = (REVERSE acc ++ FILTER (λ(c,x). f x ≠ NONE) l, + k + SUM + (MAP + (λ(c,x). if 0 ≤ c ⇔ OUTL (THE (f x)) then Num (ABS c) else 0) + (FILTER (λ(c,x). f x ≠ NONE) l))) +Proof + Induct>>simp[obj_single_aux_def,FORALL_PROD]>>rw[]>> + gvs[AllCaseEqs()]>> + first_x_assum drule>> + simp[]>>rw[] +QED + +Theorem obj_single_eq: + obj_single f l = SOME res ⇒ + obj_constraint f (l,b) = res +Proof + Cases_on`l`>>simp[obj_single_def]>>rw[] + >- + EVAL_TAC>> + Cases_on`h`>>fs[obj_single_def] >> + gvs[AllCaseEqs()]>> + drule obj_single_aux_eq_SOME>>rw[]>> + DEP_REWRITE_TAC[obj_constraint_simp]>> + simp[] +QED + +Definition full_obj_single_def: + full_obj_single f lb = + case obj_single f (FST lb) of + NONE => obj_constraint f lb + | SOME res => res +End + +(* DO NOT TRANSLATE DIRECTLY *) +Theorem obj_constraint_rewrite: + full_obj_single f lb = + obj_constraint f lb +Proof + rw[full_obj_single_def]>> + every_case_tac>>simp[]>> + drule obj_single_eq>> + metis_tac[PAIR] +QED + (* Fast substitution for obj_constraint if it is in vomap *) Definition fast_obj_constraint_def: fast_obj_constraint s l vomap = case s of INR v => if length v = 0 then ([],0) - else obj_constraint (subst_fun s) l + else full_obj_single (subst_fun s) l | INL (n,_) => if n < strlen vomap then if strsub vomap n = ^zw then ([],0) else - obj_constraint (subst_fun s) l + full_obj_single (subst_fun s) l else ([],0) End @@ -1609,7 +1872,7 @@ Theorem vomap_rel_fast_obj_constraint: fast_obj_constraint s l vomap = obj_constraint (subst_fun s) l Proof - rw[fast_obj_constraint_def]>> + rw[fast_obj_constraint_def,obj_constraint_rewrite]>> every_case_tac>> Cases_on`l`>> fs[npbcTheory.obj_constraint_def,subst_fun_def]>> From b31babd207616ad1ee79130bae459fcf4db79fb9 Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sun, 4 Feb 2024 19:08:20 +0800 Subject: [PATCH 34/38] inline map snd for obj --- .../array/npbc_arrayProgScript.sml | 22 ++++++++++--------- 1 file changed, 12 insertions(+), 10 deletions(-) diff --git a/examples/pseudo_bool/array/npbc_arrayProgScript.sml b/examples/pseudo_bool/array/npbc_arrayProgScript.sml index 5317733473..0063d8802c 100644 --- a/examples/pseudo_bool/array/npbc_arrayProgScript.sml +++ b/examples/pseudo_bool/array/npbc_arrayProgScript.sml @@ -1,7 +1,7 @@ (* Refine npbc_list to npbc_array *) -open preamble basis npbcTheory npbc_listTheory; +open preamble basis npbcTheory npbc_listTheory fastbuild; val _ = new_theory "npbc_arrayProg" @@ -3254,7 +3254,8 @@ val fold_update_resize_bitset = process_topdecs` fun fold_update_resize_bitset ls acc = case ls of [] => acc - | (x::xs) => + | (cx::xs) => + case cx of (c,x) => if x < Word8Array.length acc then (Word8Array.update acc x w8o; @@ -3270,20 +3271,22 @@ val fold_update_resize_bitset = process_topdecs` Theorem fold_update_resize_bitset_spec: ∀ls lsv accv accls. - LIST_TYPE NUM ls lsv + LIST_TYPE (PAIR_TYPE INT NUM) ls lsv ⇒ app (p : 'ffi ffi_proj) ^(fetch_v "fold_update_resize_bitset" (get_ml_prog_state())) [lsv; accv] (W8ARRAY accv accls) (POSTv v. - W8ARRAY v (FOLDL (λacc i. update_resize acc w8z w8o i) accls ls)) + W8ARRAY v (FOLDL (λacc (c,i). update_resize acc w8z w8o i) accls ls)) Proof Induct>> xcf "fold_update_resize_bitset" (get_ml_prog_state ())>> gvs[LIST_TYPE_def]>>xmatch >- ( xvar>>xsimpl)>> + pairarg_tac>>fs[PAIR_TYPE_def]>> + xmatch>> assume_tac w8o_v_thm>> assume_tac w8z_v_thm>> rpt xlet_autop>> @@ -3301,8 +3304,7 @@ val mk_vomap_arr = process_topdecs` fun mk_vomap_arr n fc = let val acc = Word8Array.array n w8z - val f = map_snd (fst fc) - val acc = fold_update_resize_bitset f acc in + val acc = fold_update_resize_bitset (fst fc) acc in Word8Array.substring acc 0 (Word8Array.length acc) end` |> append_prog; @@ -3310,12 +3312,12 @@ Theorem map_foldl_rel: ∀ls accA accB. MAP (CHR o w2n) accA = accB ⇒ MAP (CHR ∘ w2n) - (FOLDL (λacc i. update_resize acc w8z w8o i) accA ls) = - FOLDL (λacc i. update_resize acc #"\^@" #"\^A" i) accB ls + (FOLDL (λacc (c,i). update_resize acc w8z w8o i) accA ls) = + FOLDL (λacc i. update_resize acc #"\^@" #"\^A" i) accB (MAP SND ls) Proof Induct>>rw[]>> first_x_assum match_mp_tac>> - rw[update_resize_def,LUPDATE_MAP]>> + Cases_on`h`>>rw[update_resize_def,LUPDATE_MAP]>> EVAL_TAC QED @@ -3343,7 +3345,7 @@ Proof qmatch_goalsub_abbrev_tac`strlit B`>> qsuff_tac`A = B`>- metis_tac[]>> unabbrev_all_tac>> - simp[map_snd_def]>> + simp[]>> match_mp_tac map_foldl_rel>> simp[map_replicate]>> EVAL_TAC From d05733f2aa3316f6e5091a4174d38516883e776e Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sun, 4 Feb 2024 21:03:44 +0800 Subject: [PATCH 35/38] remove stray fastbuild --- examples/pseudo_bool/array/npbc_arrayProgScript.sml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/examples/pseudo_bool/array/npbc_arrayProgScript.sml b/examples/pseudo_bool/array/npbc_arrayProgScript.sml index 0063d8802c..9246ca1fe9 100644 --- a/examples/pseudo_bool/array/npbc_arrayProgScript.sml +++ b/examples/pseudo_bool/array/npbc_arrayProgScript.sml @@ -1,7 +1,7 @@ (* Refine npbc_list to npbc_array *) -open preamble basis npbcTheory npbc_listTheory fastbuild; +open preamble basis npbcTheory npbc_listTheory; val _ = new_theory "npbc_arrayProg" From 5e24a97aa99713e8705714adcc4277d84ef754ca Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Sun, 4 Feb 2024 23:45:24 +0800 Subject: [PATCH 36/38] eta expand subst_fun --- .../pseudo_bool/array/npbc_arrayProgScript.sml | 16 +++++++++++++++- 1 file changed, 15 insertions(+), 1 deletion(-) diff --git a/examples/pseudo_bool/array/npbc_arrayProgScript.sml b/examples/pseudo_bool/array/npbc_arrayProgScript.sml index 9246ca1fe9..04b97569c9 100644 --- a/examples/pseudo_bool/array/npbc_arrayProgScript.sml +++ b/examples/pseudo_bool/array/npbc_arrayProgScript.sml @@ -1217,7 +1217,21 @@ val res = translate subst_lhs_def; val res = translate subst_def; val res = translate obj_constraint_def; -val res = translate npbc_checkTheory.subst_fun_def; + +Theorem subst_fun_alt: + subst_fun (s:subst) = + case s of + INL (m,v) => \n. if n = m then SOME v else NONE + | INR s => + \n. if n < length s then sub s n else NONE +Proof + rw[FUN_EQ_THM]>> + every_case_tac>> + EVAL_TAC>> + rw[] +QED + +val res = translate subst_fun_alt; val res = translate subst_subst_fun_def; val extract_clauses_arr = process_topdecs` From b7a586017bebc526feee6764a1af42a8f21015b6 Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Mon, 5 Feb 2024 11:17:42 +0800 Subject: [PATCH 37/38] add simplified wcnfProof thm --- .../compilation/proofs/wcnfProofScript.sml | 92 +++++++++++++++++++ 1 file changed, 92 insertions(+) diff --git a/examples/pseudo_bool/array/compilation/proofs/wcnfProofScript.sml b/examples/pseudo_bool/array/compilation/proofs/wcnfProofScript.sml index 19ee1c7a35..513b1e3ddf 100644 --- a/examples/pseudo_bool/array/compilation/proofs/wcnfProofScript.sml +++ b/examples/pseudo_bool/array/compilation/proofs/wcnfProofScript.sml @@ -124,4 +124,96 @@ QED val chk = machine_code_sound |> check_thm; +(* Showing how to specialize the main theorem *) + +Theorem strcat_cancel: + a = b ∧ a ^ y = b ^ z ⇒ + y = z +Proof + EVAL_TAC>>rw[]>> + every_case_tac>>fs[] +QED + +Theorem isSuffix_STRCAT: + isSuffix x (y ^ x) +Proof + rw[mlstringTheory.isSuffix_def]>> + Cases_on`x`>>Cases_on`y`>> + simp[mlstringTheory.strlit_STRCAT]>> + DEP_REWRITE_TAC[mlstringTheory.isStringThere_SEG]>> + rw[]>> + gvs[SEG_TAKE_DROP]>> + simp[DROP_LENGTH_APPEND] +QED + +Theorem strcat_isSuffix: + x ^ y = a ^ b ⇒ + isSuffix y b ∨ isSuffix b y +Proof + map_every Cases_on [`x`,`y`,`a`,`b`]>> + simp[mlstringTheory.strlit_STRCAT,listTheory.APPEND_EQ_APPEND]>>rw[]>> + metis_tac[isSuffix_STRCAT,mlstringTheory.strlit_STRCAT] +QED + +Theorem strcat_isSuffixL: + x ^ y = a ⇒ + isSuffix y a +Proof + `a = a ^ strlit ""` by + (Cases_on`a`>>EVAL_TAC)>> + rw[]>> + metis_tac[isSuffix_STRCAT] +QED + +Theorem isSuffix_exists: + isSuffix x y ⇒ + ∃z. y = z ^ x +Proof + rw[mlstringTheory.isSuffix_def]>> + Cases_on`x`>>Cases_on`y`>> + pop_assum mp_tac>> + DEP_REWRITE_TAC[mlstringTheory.isStringThere_SEG]>> + rw[]>> + gvs[SEG_TAKE_DROP]>> + qexists_tac`strlit (TAKE (STRLEN s' − STRLEN s) s')`>> + simp[mlstringTheory.strlit_STRCAT]>> + rename1`l ≤ STRLEN s'`>> + pop_assum SUBST_ALL_TAC>>simp[GSYM TAKE_SUM] +QED + +Theorem machine_code_sound_equiopt: + cake_pb_wcnf_run cl fs mc ms ⇒ + ∃out err. + extract_fs fs (cake_pb_wcnf_io_events cl fs) = + SOME (add_stdout (add_stderr fs err) out) ∧ + ( + LENGTH cl = 4 ∧ + isSuffix «s VERIFIED OUTPUT EQUIOPTIMAL\n» out ⇒ + ∃wfml wfml'. + inFS_fname fs (EL 1 cl) ∧ inFS_fname fs (EL 3 cl) ∧ + get_fml fs (EL 1 cl) = SOME wfml ∧ + get_fml fs (EL 3 cl) = SOME wfml' ∧ + opt_cost wfml = opt_cost wfml' + ) +Proof + rw[]>> + drule machine_code_sound>>rw[]>> + first_x_assum (irule_at Any)>> + rw[]>>drule isSuffix_exists>> + pop_assum kall_tac>> + strip_tac>> + rename1`pref ^ _`>> + gvs[]>> + `pref ^ «s VERIFIED OUTPUT EQUIOPTIMAL\n» ≠ «»` by + (EVAL_TAC>> + rw[])>> + gvs[]>> + `iseqopt` by + (drule strcat_isSuffix>> + simp[wcnfProgTheory.print_maxsat_output_str_def]>> + IF_CASES_TAC>>simp[]>> + EVAL_TAC)>> + gvs[wcnfProgTheory.maxsat_output_sem_def] +QED + val _ = export_theory(); From 9bc161d6b777d41e1956cbcd830ee536e410c65c Mon Sep 17 00:00:00 2001 From: tanyongkiam Date: Wed, 14 Feb 2024 19:12:19 +0800 Subject: [PATCH 38/38] add -j1 to Holmakefile --- examples/pseudo_bool/array/compilation/Holmakefile | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/examples/pseudo_bool/array/compilation/Holmakefile b/examples/pseudo_bool/array/compilation/Holmakefile index b4395f17a5..9a1f705151 100644 --- a/examples/pseudo_bool/array/compilation/Holmakefile +++ b/examples/pseudo_bool/array/compilation/Holmakefile @@ -1,5 +1,5 @@ -INCLUDES = $(CAKEMLDIR)/misc $(CAKEMLDIR)/basis $(CAKEMLDIR)/compiler .. -CLINE_OPTIONS = +INCLUDES = $(CAKEMLDIR)/misc $(CAKEMLDIR)/basis $(CAKEMLDIR)/compiler .. $(CAKEMLDIR)/unverified/sexpr-bootstrap $(CAKEMLDIR)/compiler/parsing +CLINE_OPTIONS = -j1 all: $(DEFAULT_TARGETS) README.md .PHONY: all