| author | wenzelm | 
| Fri, 06 Mar 2015 21:20:30 +0100 | |
| changeset 59627 | bb1e4a35d506 | 
| parent 59621 | 291934bac95e | 
| child 60781 | 2da59cdf531c | 
| 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: 
54406 
diff
changeset
 | 
9  | 
val pat_completeness_tac: Proof.context -> int -> tactic  | 
| 
 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 
blanchet 
parents: 
54406 
diff
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: 
54406 
diff
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: 
33083 
diff
changeset
 | 
17  | 
open Function_Lib  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents: 
33083 
diff
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  | 
|
| 59627 | 26  | 
fun inst_case_thm ctxt x P thm =  | 
| 59582 | 27  | 
let val [Pv, xv] = Term.add_vars (Thm.prop_of thm) []  | 
| 33083 | 28  | 
in  | 
| 59627 | 29  | 
thm |> cterm_instantiate (map (apply2 (Thm.cterm_of ctxt)) [(Var xv, x), (Var Pv, P)])  | 
| 33083 | 30  | 
end  | 
31  | 
||
32  | 
fun invent_vars constr i =  | 
|
33  | 
let  | 
|
34  | 
val Ts = binder_types (fastype_of constr)  | 
|
35  | 
val j = i + length Ts  | 
|
36  | 
val is = i upto (j - 1)  | 
|
37  | 
val avs = map2 mk_argvar is Ts  | 
|
38  | 
val pvs = map2 mk_patvar is Ts  | 
|
39  | 
in  | 
|
40  | 
(avs, pvs, j)  | 
|
41  | 
end  | 
|
42  | 
||
| 59627 | 43  | 
fun filter_pats _ _ _ [] = []  | 
44  | 
| filter_pats _ _ _ (([], _) :: _) = raise Match  | 
|
| 59618 | 45  | 
| filter_pats ctxt cons pvars (((pat as Free _) :: pats, thm) :: pts) =  | 
| 59627 | 46  | 
let val inst = list_comb (cons, pvars) in  | 
47  | 
(inst :: pats, inst_free (Thm.cterm_of ctxt pat) (Thm.cterm_of ctxt inst) thm) ::  | 
|
48  | 
filter_pats ctxt cons pvars pts  | 
|
49  | 
end  | 
|
50  | 
| filter_pats ctxt cons pvars ((pat :: pats, thm) :: pts) =  | 
|
51  | 
if fst (strip_comb pat) = cons  | 
|
52  | 
then (pat :: pats, thm) :: filter_pats ctxt cons pvars pts  | 
|
53  | 
else filter_pats ctxt cons pvars pts  | 
|
| 33083 | 54  | 
|
55  | 
||
| 59627 | 56  | 
fun transform_pat _ _ _ ([] , _) = raise Match  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
47432 
diff
changeset
 | 
57  | 
| transform_pat ctxt avars c_assum (pat :: pats, thm) =  | 
| 59618 | 58  | 
let  | 
59  | 
val (_, subps) = strip_comb pat  | 
|
| 59627 | 60  | 
val eqs = map (Thm.cterm_of ctxt o HOLogic.mk_Trueprop o HOLogic.mk_eq) (avars ~~ subps)  | 
| 59618 | 61  | 
val c_eq_pat =  | 
62  | 
simplify (put_simpset HOL_basic_ss ctxt addsimps (map Thm.assume eqs)) c_assum  | 
|
63  | 
in  | 
|
64  | 
(subps @ pats,  | 
|
65  | 
fold_rev Thm.implies_intr eqs (Thm.implies_elim thm c_eq_pat))  | 
|
66  | 
end  | 
|
| 33083 | 67  | 
|
68  | 
||
69  | 
exception COMPLETENESS  | 
|
70  | 
||
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
47432 
diff
changeset
 | 
71  | 
fun constr_case ctxt P idx (v :: vs) pats cons =  | 
| 59618 | 72  | 
let  | 
73  | 
val (avars, pvars, newidx) = invent_vars cons idx  | 
|
74  | 
val c_hyp =  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
75  | 
Thm.cterm_of ctxt  | 
| 59618 | 76  | 
(HOLogic.mk_Trueprop (HOLogic.mk_eq (v, list_comb (cons, avars))))  | 
77  | 
val c_assum = Thm.assume c_hyp  | 
|
78  | 
val newpats = map (transform_pat ctxt avars c_assum) (filter_pats ctxt cons pvars pats)  | 
|
79  | 
in  | 
|
80  | 
o_alg ctxt P newidx (avars @ vs) newpats  | 
|
81  | 
|> Thm.implies_intr c_hyp  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
82  | 
|> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt) avars  | 
| 59618 | 83  | 
end  | 
| 33083 | 84  | 
| constr_case _ _ _ _ _ _ = raise Match  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
47432 
diff
changeset
 | 
85  | 
and o_alg _ P idx [] (([], Pthm) :: _) = Pthm  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
47432 
diff
changeset
 | 
86  | 
| o_alg _ P idx (v :: vs) [] = raise COMPLETENESS  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
47432 
diff
changeset
 | 
87  | 
| o_alg ctxt P idx (v :: vs) pts =  | 
| 59618 | 88  | 
if forall (is_Free o hd o fst) pts (* Var case *)  | 
89  | 
then o_alg ctxt P idx vs  | 
|
90  | 
(map (fn (pv :: pats, thm) =>  | 
|
91  | 
(pats, refl RS  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
92  | 
(inst_free (Thm.cterm_of ctxt pv)  | 
| 
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
93  | 
(Thm.cterm_of ctxt v) thm))) pts)  | 
| 59618 | 94  | 
else (* Cons case *)  | 
95  | 
let  | 
|
96  | 
val T as Type (tname, _) = fastype_of v  | 
|
97  | 
          val SOME {exhaust=case_thm, ...} = Ctr_Sugar.ctr_sugar_of ctxt tname
 | 
|
98  | 
val constrs = inst_constrs_of ctxt T  | 
|
99  | 
val c_cases = map (constr_case ctxt P idx (v :: vs) pts) constrs  | 
|
100  | 
in  | 
|
| 59627 | 101  | 
inst_case_thm ctxt v P case_thm  | 
| 59618 | 102  | 
|> fold (curry op COMP) c_cases  | 
103  | 
end  | 
|
| 33083 | 104  | 
| o_alg _ _ _ _ _ = raise Match  | 
105  | 
||
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
47432 
diff
changeset
 | 
106  | 
fun prove_completeness ctxt xs P qss patss =  | 
| 33083 | 107  | 
let  | 
108  | 
fun mk_assum qs pats =  | 
|
109  | 
HOLogic.mk_Trueprop P  | 
|
110  | 
|> fold_rev (curry Logic.mk_implies o HOLogic.mk_Trueprop o HOLogic.mk_eq) (xs ~~ pats)  | 
|
111  | 
|> fold_rev Logic.all qs  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
112  | 
|> Thm.cterm_of ctxt  | 
| 33083 | 113  | 
|
114  | 
val hyps = map2 mk_assum qss patss  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
115  | 
fun inst_hyps hyp qs = fold (Thm.forall_elim o Thm.cterm_of ctxt) qs (Thm.assume hyp)  | 
| 33083 | 116  | 
val assums = map2 inst_hyps hyps qss  | 
117  | 
in  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
47432 
diff
changeset
 | 
118  | 
o_alg ctxt P 2 xs (patss ~~ assums)  | 
| 36945 | 119  | 
|> fold_rev Thm.implies_intr hyps  | 
| 33083 | 120  | 
end  | 
121  | 
||
122  | 
fun pat_completeness_tac ctxt = SUBGOAL (fn (subgoal, i) =>  | 
|
123  | 
let  | 
|
124  | 
val (vs, subgf) = dest_all_all subgoal  | 
|
125  | 
val (cases, _ $ thesis) = Logic.strip_horn subgf  | 
|
126  | 
handle Bind => raise COMPLETENESS  | 
|
127  | 
||
128  | 
fun pat_of assum =  | 
|
129  | 
let  | 
|
130  | 
val (qs, imp) = dest_all_all assum  | 
|
131  | 
val prems = Logic.strip_imp_prems imp  | 
|
132  | 
in  | 
|
133  | 
(qs, map (HOLogic.dest_eq o HOLogic.dest_Trueprop) prems)  | 
|
134  | 
end  | 
|
135  | 
||
136  | 
val (qss, x_pats) = split_list (map pat_of cases)  | 
|
137  | 
val xs = map fst (hd x_pats)  | 
|
| 
47060
 
e2741ec9ae36
prefer explicitly qualified exception List.Empty;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
138  | 
handle List.Empty => raise COMPLETENESS  | 
| 33083 | 139  | 
|
140  | 
val patss = map (map snd) x_pats  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
47432 
diff
changeset
 | 
141  | 
val complete_thm = prove_completeness ctxt xs thesis qss patss  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
142  | 
|> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt) vs  | 
| 33083 | 143  | 
in  | 
| 52467 | 144  | 
PRIMITIVE (fn st => Drule.compose (complete_thm, i, st))  | 
| 33083 | 145  | 
end  | 
146  | 
handle COMPLETENESS => no_tac)  | 
|
147  | 
||
148  | 
end  |