| author | wenzelm | 
| Sun, 12 Jan 2014 14:32:22 +0100 | |
| changeset 54998 | 8601434fa334 | 
| parent 54407 | e95831757903 | 
| child 59058 | a78612c67ec0 | 
| permissions | -rw-r--r-- | 
| 37744 | 1 | (* Title: HOL/Tools/Function/pat_completeness.ML | 
| 33083 | 2 | Author: Alexander Krauss, TU Muenchen | 
| 3 | ||
| 4 | Method "pat_completeness" to prove completeness of datatype patterns. | |
| 5 | *) | |
| 6 | ||
| 7 | signature PAT_COMPLETENESS = | |
| 8 | sig | |
| 54407 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
54406diff
changeset | 9 | val pat_completeness_tac: Proof.context -> int -> tactic | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
54406diff
changeset | 10 | val prove_completeness: Proof.context -> term list -> term -> term list list -> | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
54406diff
changeset | 11 | term list list -> thm | 
| 33083 | 12 | end | 
| 13 | ||
| 14 | structure Pat_Completeness : PAT_COMPLETENESS = | |
| 15 | struct | |
| 16 | ||
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33083diff
changeset | 17 | open Function_Lib | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33083diff
changeset | 18 | open Function_Common | 
| 33083 | 19 | |
| 20 | ||
| 21 | fun mk_argvar i T = Free ("_av" ^ (string_of_int i), T)
 | |
| 22 | fun mk_patvar i T = Free ("_pv" ^ (string_of_int i), T)
 | |
| 23 | ||
| 36945 | 24 | fun inst_free var inst = Thm.forall_elim inst o Thm.forall_intr var | 
| 33083 | 25 | |
| 26 | fun inst_case_thm thy x P thm = | |
| 27 | let val [Pv, xv] = Term.add_vars (prop_of thm) [] | |
| 28 | in | |
| 29 | thm |> cterm_instantiate (map (pairself (cterm_of thy)) | |
| 30 | [(Var xv, x), (Var Pv, P)]) | |
| 31 | end | |
| 32 | ||
| 33 | fun invent_vars constr i = | |
| 34 | let | |
| 35 | val Ts = binder_types (fastype_of constr) | |
| 36 | val j = i + length Ts | |
| 37 | val is = i upto (j - 1) | |
| 38 | val avs = map2 mk_argvar is Ts | |
| 39 | val pvs = map2 mk_patvar is Ts | |
| 40 | in | |
| 41 | (avs, pvs, j) | |
| 42 | end | |
| 43 | ||
| 44 | fun filter_pats thy cons pvars [] = [] | |
| 45 | | filter_pats thy cons pvars (([], thm) :: pts) = raise Match | |
| 46 | | filter_pats thy cons pvars (((pat as Free _) :: pats, thm) :: pts) = | |
| 47 | let val inst = list_comb (cons, pvars) | |
| 48 | in (inst :: pats, inst_free (cterm_of thy pat) (cterm_of thy inst) thm) | |
| 49 | :: (filter_pats thy cons pvars pts) | |
| 50 | end | |
| 51 | | filter_pats thy cons pvars ((pat :: pats, thm) :: pts) = | |
| 52 | if fst (strip_comb pat) = cons | |
| 53 | then (pat :: pats, thm) :: (filter_pats thy cons pvars pts) | |
| 54 | else filter_pats thy cons pvars pts | |
| 55 | ||
| 56 | ||
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 57 | fun transform_pat _ avars c_assum ([] , thm) = raise Match | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 58 | | transform_pat ctxt avars c_assum (pat :: pats, thm) = | 
| 33083 | 59 | let | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 60 | val thy = Proof_Context.theory_of ctxt | 
| 33083 | 61 | val (_, subps) = strip_comb pat | 
| 62 | val eqs = map (cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (avars ~~ subps) | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 63 | val c_eq_pat = simplify (put_simpset HOL_basic_ss ctxt addsimps (map Thm.assume eqs)) c_assum | 
| 33083 | 64 | in | 
| 65 | (subps @ pats, | |
| 36945 | 66 | fold_rev Thm.implies_intr eqs (Thm.implies_elim thm c_eq_pat)) | 
| 33083 | 67 | end | 
| 68 | ||
| 69 | ||
| 70 | exception COMPLETENESS | |
| 71 | ||
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 72 | fun constr_case ctxt P idx (v :: vs) pats cons = | 
| 33083 | 73 | let | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 74 | val thy = Proof_Context.theory_of ctxt | 
| 33083 | 75 | val (avars, pvars, newidx) = invent_vars cons idx | 
| 76 | val c_hyp = cterm_of thy (HOLogic.mk_Trueprop (HOLogic.mk_eq (v, list_comb (cons, avars)))) | |
| 36945 | 77 | val c_assum = Thm.assume c_hyp | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 78 | val newpats = map (transform_pat ctxt avars c_assum) (filter_pats thy cons pvars pats) | 
| 33083 | 79 | in | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 80 | o_alg ctxt P newidx (avars @ vs) newpats | 
| 36945 | 81 | |> Thm.implies_intr c_hyp | 
| 82 | |> fold_rev (Thm.forall_intr o cterm_of thy) avars | |
| 33083 | 83 | end | 
| 84 | | constr_case _ _ _ _ _ _ = raise Match | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 85 | and o_alg _ P idx [] (([], Pthm) :: _) = Pthm | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 86 | | o_alg _ P idx (v :: vs) [] = raise COMPLETENESS | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 87 | | o_alg ctxt P idx (v :: vs) pts = | 
| 33083 | 88 | if forall (is_Free o hd o fst) pts (* Var case *) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 89 | then o_alg ctxt P idx vs | 
| 33083 | 90 | (map (fn (pv :: pats, thm) => | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 91 | (pats, refl RS | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 92 | (inst_free (cterm_of (Proof_Context.theory_of ctxt) pv) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 93 | (cterm_of (Proof_Context.theory_of ctxt) v) thm))) pts) | 
| 33083 | 94 | else (* Cons case *) | 
| 95 | let | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 96 | val thy = Proof_Context.theory_of ctxt | 
| 54407 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
54406diff
changeset | 97 | val T as Type (tname, _) = fastype_of v | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
54406diff
changeset | 98 |       val SOME {exhaust=case_thm, ...} = Ctr_Sugar.ctr_sugar_of ctxt tname
 | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
54406diff
changeset | 99 | val constrs = inst_constrs_of ctxt T | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 100 | val c_cases = map (constr_case ctxt P idx (v :: vs) pts) constrs | 
| 33083 | 101 | in | 
| 102 | inst_case_thm thy v P case_thm | |
| 103 | |> fold (curry op COMP) c_cases | |
| 104 | end | |
| 105 | | o_alg _ _ _ _ _ = raise Match | |
| 106 | ||
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 107 | fun prove_completeness ctxt xs P qss patss = | 
| 33083 | 108 | let | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 109 | val thy = Proof_Context.theory_of ctxt | 
| 33083 | 110 | fun mk_assum qs pats = | 
| 111 | HOLogic.mk_Trueprop P | |
| 112 | |> fold_rev (curry Logic.mk_implies o HOLogic.mk_Trueprop o HOLogic.mk_eq) (xs ~~ pats) | |
| 113 | |> fold_rev Logic.all qs | |
| 114 | |> cterm_of thy | |
| 115 | ||
| 116 | val hyps = map2 mk_assum qss patss | |
| 36945 | 117 | fun inst_hyps hyp qs = fold (Thm.forall_elim o cterm_of thy) qs (Thm.assume hyp) | 
| 33083 | 118 | val assums = map2 inst_hyps hyps qss | 
| 119 | in | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 120 | o_alg ctxt P 2 xs (patss ~~ assums) | 
| 36945 | 121 | |> fold_rev Thm.implies_intr hyps | 
| 33083 | 122 | end | 
| 123 | ||
| 124 | fun pat_completeness_tac ctxt = SUBGOAL (fn (subgoal, i) => | |
| 125 | let | |
| 42361 | 126 | val thy = Proof_Context.theory_of ctxt | 
| 33083 | 127 | val (vs, subgf) = dest_all_all subgoal | 
| 128 | val (cases, _ $ thesis) = Logic.strip_horn subgf | |
| 129 | handle Bind => raise COMPLETENESS | |
| 130 | ||
| 131 | fun pat_of assum = | |
| 132 | let | |
| 133 | val (qs, imp) = dest_all_all assum | |
| 134 | val prems = Logic.strip_imp_prems imp | |
| 135 | in | |
| 136 | (qs, map (HOLogic.dest_eq o HOLogic.dest_Trueprop) prems) | |
| 137 | end | |
| 138 | ||
| 139 | val (qss, x_pats) = split_list (map pat_of cases) | |
| 140 | val xs = map fst (hd x_pats) | |
| 47060 
e2741ec9ae36
prefer explicitly qualified exception List.Empty;
 wenzelm parents: 
42361diff
changeset | 141 | handle List.Empty => raise COMPLETENESS | 
| 33083 | 142 | |
| 143 | val patss = map (map snd) x_pats | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47432diff
changeset | 144 | val complete_thm = prove_completeness ctxt xs thesis qss patss | 
| 36945 | 145 | |> fold_rev (Thm.forall_intr o cterm_of thy) vs | 
| 33083 | 146 | in | 
| 52467 | 147 | PRIMITIVE (fn st => Drule.compose (complete_thm, i, st)) | 
| 33083 | 148 | end | 
| 149 | handle COMPLETENESS => no_tac) | |
| 150 | ||
| 151 | end |