| author | wenzelm | 
| Thu, 30 Mar 2023 11:58:53 +0200 | |
| changeset 77756 | efd5c582d7ae | 
| parent 69593 | 3dda49e08b9d | 
| child 81954 | 6f2bcdfa9a19 | 
| permissions | -rw-r--r-- | 
| 55061 | 1  | 
(* Title: HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML  | 
| 
49123
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents:  
diff
changeset
 | 
2  | 
Author: Jasmin Blanchette, TU Muenchen  | 
| 57668 | 3  | 
Author: Martin Desharnais, TU Muenchen  | 
4  | 
Copyright 2012, 2013, 2014  | 
|
| 
49123
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents:  
diff
changeset
 | 
5  | 
|
| 49389 | 6  | 
Tactics for datatype and codatatype sugar.  | 
| 
49123
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents:  
diff
changeset
 | 
7  | 
*)  | 
| 
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents:  
diff
changeset
 | 
8  | 
|
| 49636 | 9  | 
signature BNF_FP_DEF_SUGAR_TACTICS =  | 
| 
49123
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents:  
diff
changeset
 | 
10  | 
sig  | 
| 63842 | 11  | 
val sumprod_thms_rel: thm list  | 
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49542 
diff
changeset
 | 
12  | 
|
| 64607 | 13  | 
val co_induct_inst_as_projs_tac: Proof.context -> int -> tactic  | 
| 58093 | 14  | 
val mk_case_transfer_tac: Proof.context -> thm -> thm list -> tactic  | 
| 64636 | 15  | 
val mk_coinduct_discharge_prem_tac: Proof.context -> thm list -> thm list -> int -> int -> int ->  | 
16  | 
thm -> thm -> thm -> thm -> thm -> thm list -> thm list list -> thm list list -> int -> tactic  | 
|
| 
49591
 
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
 
blanchet 
parents: 
49590 
diff
changeset
 | 
17  | 
val mk_coinduct_tac: Proof.context -> thm list -> int -> int list -> thm -> thm list ->  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
18  | 
thm list -> thm list -> thm list -> thm list -> thm list list -> thm list list list ->  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
19  | 
thm list list list -> tactic  | 
| 55867 | 20  | 
val mk_corec_tac: thm list -> thm list -> thm -> thm -> thm -> thm -> Proof.context -> tactic  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57893 
diff
changeset
 | 
21  | 
val mk_corec_disc_iff_tac: thm list -> thm list -> thm list -> Proof.context -> tactic  | 
| 58734 | 22  | 
val mk_co_rec_o_map_tac: Proof.context -> thm -> thm list -> thm list -> thm list -> thm -> thm ->  | 
23  | 
thm Seq.seq  | 
|
| 58448 | 24  | 
val mk_corec_transfer_tac: Proof.context -> cterm list -> cterm list -> thm list -> thm list ->  | 
25  | 
thm list -> thm list -> thm list -> ''a list -> ''a list list -> ''a list list list list ->  | 
|
26  | 
''a list list list list -> tactic  | 
|
| 49501 | 27  | 
val mk_ctor_iff_dtor_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm ->  | 
28  | 
tactic  | 
|
| 58327 | 29  | 
val mk_ctr_transfer_tac: Proof.context -> thm list -> thm list -> tactic  | 
| 60728 | 30  | 
val mk_disc_transfer_tac: Proof.context -> thm -> thm -> thm list -> tactic  | 
| 
49161
 
a8e74375d971
fixed (n + 1)st bug in "mk_exhaust_tac" -- arose with uncurried constructors
 
blanchet 
parents: 
49160 
diff
changeset
 | 
31  | 
val mk_exhaust_tac: Proof.context -> int -> thm list -> thm -> thm -> tactic  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
32  | 
val mk_half_distinct_tac: Proof.context -> thm -> thm -> thm list -> tactic  | 
| 
64628
 
19bc22274cd9
export ML function (towards nonuniform datatypes)
 
blanchet 
parents: 
64624 
diff
changeset
 | 
33  | 
val mk_induct_discharge_prem_tac: Proof.context -> int -> int -> thm list -> thm list ->  | 
| 
 
19bc22274cd9
export ML function (towards nonuniform datatypes)
 
blanchet 
parents: 
64624 
diff
changeset
 | 
34  | 
thm list -> thm list -> int -> int -> int list -> tactic  | 
| 49590 | 35  | 
val mk_induct_tac: Proof.context -> int -> int list -> int list list -> int list list list ->  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
36  | 
thm list -> thm -> thm list -> thm list -> thm list -> thm list list -> tactic  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
37  | 
val mk_inject_tac: Proof.context -> thm -> thm -> thm -> tactic  | 
| 
64415
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64067 
diff
changeset
 | 
38  | 
val mk_map_tac: Proof.context -> thm list -> thm -> thm -> thm list -> thm list -> thm list ->  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64067 
diff
changeset
 | 
39  | 
tactic  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57893 
diff
changeset
 | 
40  | 
val mk_map_disc_iff_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> tactic  | 
| 
58326
 
7e142efcee1a
make 'rel_sel' and 'map_sel' tactics more robust
 
desharna 
parents: 
58181 
diff
changeset
 | 
41  | 
val mk_map_sel_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> thm list ->  | 
| 
 
7e142efcee1a
make 'rel_sel' and 'map_sel' tactics more robust
 
desharna 
parents: 
58181 
diff
changeset
 | 
42  | 
thm list -> tactic  | 
| 55867 | 43  | 
val mk_rec_tac: thm list -> thm list -> thm list -> thm -> thm -> thm -> thm -> Proof.context ->  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
44  | 
tactic  | 
| 58446 | 45  | 
val mk_rec_transfer_tac: Proof.context -> int -> int list -> cterm list -> cterm list ->  | 
| 58966 | 46  | 
term list list list list -> thm list -> thm list -> thm list -> thm list -> tactic  | 
| 
64415
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64067 
diff
changeset
 | 
47  | 
val mk_rel_tac: Proof.context -> thm list -> thm -> thm -> thm list -> thm list -> thm list ->  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64067 
diff
changeset
 | 
48  | 
tactic  | 
| 62335 | 49  | 
val mk_rel_case_tac: Proof.context -> cterm -> cterm -> thm -> thm list -> thm list -> thm list ->  | 
50  | 
thm list -> thm list -> tactic  | 
|
| 
57301
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57152 
diff
changeset
 | 
51  | 
val mk_rel_coinduct0_tac: Proof.context -> thm -> cterm list -> thm list -> thm list ->  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57152 
diff
changeset
 | 
52  | 
thm list list -> thm list list -> thm list list -> thm list -> thm list -> thm list ->  | 
| 57471 | 53  | 
thm list -> thm list -> thm list -> tactic  | 
54  | 
val mk_rel_induct0_tac: Proof.context -> thm -> thm list -> cterm list -> thm list ->  | 
|
55  | 
thm list list -> thm list -> thm list -> thm list -> thm list -> tactic  | 
|
| 57563 | 56  | 
val mk_rel_sel_tac: Proof.context -> cterm -> cterm -> thm -> thm list -> thm list -> thm list ->  | 
| 
58326
 
7e142efcee1a
make 'rel_sel' and 'map_sel' tactics more robust
 
desharna 
parents: 
58181 
diff
changeset
 | 
57  | 
thm list -> thm list -> thm list -> tactic  | 
| 58676 | 58  | 
val mk_sel_transfer_tac: Proof.context -> int -> thm list -> thm -> tactic  | 
| 63851 | 59  | 
val mk_set0_tac: Proof.context -> thm list -> thm list -> thm -> thm list -> thm list ->  | 
60  | 
thm list -> thm list -> thm list -> tactic  | 
|
| 57893 | 61  | 
val mk_set_cases_tac: Proof.context -> cterm -> thm list -> thm -> thm list -> tactic  | 
| 57700 | 62  | 
val mk_set_induct0_tac: Proof.context -> cterm list -> thm list -> thm list -> thm list ->  | 
63  | 
thm list -> thm list -> thm list -> thm list -> tactic  | 
|
| 57891 | 64  | 
val mk_set_intros_tac: Proof.context -> thm list -> tactic  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57893 
diff
changeset
 | 
65  | 
val mk_set_sel_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> thm list -> tactic  | 
| 
49123
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents:  
diff
changeset
 | 
66  | 
end;  | 
| 
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents:  
diff
changeset
 | 
67  | 
|
| 49636 | 68  | 
structure BNF_FP_Def_Sugar_Tactics : BNF_FP_DEF_SUGAR_TACTICS =  | 
| 
49123
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents:  
diff
changeset
 | 
69  | 
struct  | 
| 
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents:  
diff
changeset
 | 
70  | 
|
| 
56991
 
8e9ca31e9b8e
generate 'disc_map_iff[simp]' theorem for (co)datatypes
 
desharna 
parents: 
56990 
diff
changeset
 | 
71  | 
open Ctr_Sugar_Util  | 
| 49125 | 72  | 
open BNF_Tactics  | 
| 
49123
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents:  
diff
changeset
 | 
73  | 
open BNF_Util  | 
| 
51850
 
106afdf5806c
renamed a few FP-related files, to make it clear that these are not the sum of LFP + GFP but rather shared basic libraries
 
blanchet 
parents: 
51843 
diff
changeset
 | 
74  | 
open BNF_FP_Util  | 
| 
49123
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents:  
diff
changeset
 | 
75  | 
|
| 58446 | 76  | 
val case_sum_transfer = @{thm case_sum_transfer};
 | 
| 67399 | 77  | 
val case_sum_transfer_eq = @{thm case_sum_transfer[of "(=)" _ "(=)", simplified sum.rel_eq]};
 | 
| 58446 | 78  | 
val case_prod_transfer = @{thm case_prod_transfer};
 | 
| 67399 | 79  | 
val case_prod_transfer_eq = @{thm case_prod_transfer[of "(=)" "(=)", simplified prod.rel_eq]};
 | 
| 58446 | 80  | 
|
| 49590 | 81  | 
val basic_simp_thms = @{thms simp_thms(7,8,12,14,22,24)};
 | 
82  | 
val more_simp_thms = basic_simp_thms @ @{thms simp_thms(11,15,16,21)};
 | 
|
| 57303 | 83  | 
val simp_thms' = @{thms simp_thms(6,7,8,11,12,15,16,22,24)};
 | 
| 49590 | 84  | 
|
| 55966 | 85  | 
val sumprod_thms_map = @{thms id_apply map_prod_simp prod.case sum.case map_sum.simps};
 | 
| 63841 | 86  | 
val sumprod_thms_rel = @{thms rel_sum_simps rel_prod_inject prod.inject id_apply conj_assoc};
 | 
| 57891 | 87  | 
val basic_sumprod_thms_set =  | 
| 63852 | 88  | 
  @{thms UN_empty UN_insert UN_iff Un_empty_left Un_empty_right Un_iff Union_Un_distrib o_apply
 | 
89  | 
map_prod_simp mem_Collect_eq prod_set_simps map_sum.simps sum_set_simps};  | 
|
| 57891 | 90  | 
val sumprod_thms_set = @{thms UN_simps(10) image_iff} @ basic_sumprod_thms_set;
 | 
| 49368 | 91  | 
|
| 
58359
 
3782c7b0d1cc
avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
 
blanchet 
parents: 
58353 
diff
changeset
 | 
92  | 
fun is_def_looping def =  | 
| 
 
3782c7b0d1cc
avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
 
blanchet 
parents: 
58353 
diff
changeset
 | 
93  | 
(case Thm.prop_of def of  | 
| 69593 | 94  | 
Const (\<^const_name>\<open>Pure.eq\<close>, _) $ lhs $ rhs => Term.exists_subterm (curry (op aconv) lhs) rhs  | 
| 
58359
 
3782c7b0d1cc
avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
 
blanchet 
parents: 
58353 
diff
changeset
 | 
95  | 
| _ => false);  | 
| 
 
3782c7b0d1cc
avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
 
blanchet 
parents: 
58353 
diff
changeset
 | 
96  | 
|
| 49668 | 97  | 
fun hhf_concl_conv cv ctxt ct =  | 
98  | 
(case Thm.term_of ct of  | 
|
| 69593 | 99  | 
Const (\<^const_name>\<open>Pure.all\<close>, _) $ Abs _ =>  | 
| 49668 | 100  | 
Conv.arg_conv (Conv.abs_conv (hhf_concl_conv cv o snd) ctxt) ct  | 
101  | 
| _ => Conv.concl_conv ~1 cv ct);  | 
|
102  | 
||
| 54922 | 103  | 
fun co_induct_inst_as_projs ctxt k thm =  | 
| 49368 | 104  | 
let  | 
| 59582 | 105  | 
val fs = Term.add_vars (Thm.prop_of thm) []  | 
| 69593 | 106  | 
|> filter (fn (_, Type (\<^type_name>\<open>fun\<close>, [_, T'])) => T' <> HOLogic.boolT | _ => false);  | 
| 60784 | 107  | 
fun mk_inst (xi, T) = (xi, Thm.cterm_of ctxt (mk_proj T (num_binder_types T) k));  | 
| 49368 | 108  | 
in  | 
| 60784 | 109  | 
infer_instantiate ctxt (map mk_inst fs) thm  | 
| 49368 | 110  | 
end;  | 
111  | 
||
| 54922 | 112  | 
val co_induct_inst_as_projs_tac = PRIMITIVE oo co_induct_inst_as_projs;  | 
| 49368 | 113  | 
|
| 62335 | 114  | 
fun mk_case_transfer_tac ctxt rel_case cases =  | 
115  | 
let val n = length (tl (Thm.prems_of rel_case)) in  | 
|
| 60728 | 116  | 
REPEAT_DETERM (HEADGOAL (rtac ctxt rel_funI)) THEN  | 
| 62335 | 117  | 
HEADGOAL (etac ctxt rel_case) THEN  | 
| 58093 | 118  | 
ALLGOALS (hyp_subst_tac ctxt) THEN  | 
119  | 
unfold_thms_tac ctxt cases THEN  | 
|
| 60728 | 120  | 
ALLGOALS (fn k => (select_prem_tac ctxt n (dtac ctxt asm_rl) k) k) THEN  | 
121  | 
ALLGOALS (REPEAT_DETERM o (rotate_tac ~1 THEN' dtac ctxt rel_funD THEN'  | 
|
| 60752 | 122  | 
(assume_tac ctxt THEN' etac ctxt thin_rl ORELSE' rtac ctxt refl)) THEN' assume_tac ctxt)  | 
| 58093 | 123  | 
end;  | 
124  | 
||
| 58327 | 125  | 
fun mk_ctr_transfer_tac ctxt rel_intros rel_eqs =  | 
| 58095 | 126  | 
HEADGOAL Goal.conjunction_tac THEN  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59270 
diff
changeset
 | 
127  | 
ALLGOALS (REPEAT o (resolve_tac ctxt (rel_funI :: rel_intros) THEN'  | 
| 62535 | 128  | 
TRY o (REPEAT_DETERM1 o (SELECT_GOAL (unfold_thms_tac ctxt rel_eqs) THEN'  | 
129  | 
(assume_tac ctxt ORELSE' hyp_subst_tac ctxt THEN' rtac ctxt refl)))));  | 
|
| 58095 | 130  | 
|
| 60728 | 131  | 
fun mk_disc_transfer_tac ctxt rel_sel exhaust_disc distinct_disc =  | 
| 58095 | 132  | 
let  | 
133  | 
fun last_disc_tac iffD =  | 
|
| 60752 | 134  | 
HEADGOAL (rtac ctxt (rotate_prems ~1 exhaust_disc) THEN' assume_tac ctxt THEN'  | 
135  | 
REPEAT_DETERM o (rotate_tac ~1 THEN' dtac ctxt (rotate_prems 1 iffD) THEN'  | 
|
136  | 
assume_tac ctxt THEN' rotate_tac ~1 THEN'  | 
|
| 60757 | 137  | 
etac ctxt (rotate_prems 1 notE) THEN' eresolve_tac ctxt distinct_disc));  | 
| 58095 | 138  | 
in  | 
139  | 
HEADGOAL Goal.conjunction_tac THEN  | 
|
| 60728 | 140  | 
REPEAT_DETERM (HEADGOAL (rtac ctxt rel_funI THEN' dtac ctxt (rel_sel RS iffD1) THEN'  | 
| 60752 | 141  | 
REPEAT_DETERM o (etac ctxt conjE) THEN' (assume_tac ctxt ORELSE' rtac ctxt iffI))) THEN  | 
| 58095 | 142  | 
TRY (last_disc_tac iffD2) THEN TRY (last_disc_tac iffD1)  | 
143  | 
end;  | 
|
144  | 
||
| 49501 | 145  | 
fun mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor sumEN' =  | 
| 60728 | 146  | 
unfold_thms_tac ctxt (ctor_iff_dtor :: ctr_defs) THEN HEADGOAL (rtac ctxt sumEN') THEN  | 
147  | 
HEADGOAL (EVERY' (maps (fn k => [select_prem_tac ctxt n (rotate_tac 1) k,  | 
|
| 60752 | 148  | 
REPEAT_DETERM o dtac ctxt meta_spec, etac ctxt meta_mp, assume_tac ctxt]) (1 upto n)));  | 
| 49125 | 149  | 
|
| 49501 | 150  | 
fun mk_ctor_iff_dtor_tac ctxt cTs cctor cdtor ctor_dtor dtor_ctor =  | 
| 60728 | 151  | 
HEADGOAL (rtac ctxt iffI THEN'  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58507 
diff
changeset
 | 
152  | 
    EVERY' (@{map 3} (fn cTs => fn cx => fn th =>
 | 
| 60801 | 153  | 
dtac ctxt (Thm.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN'  | 
| 52324 | 154  | 
SELECT_GOAL (unfold_thms_tac ctxt [th]) THEN'  | 
| 60752 | 155  | 
assume_tac ctxt) [rev cTs, cTs] [cdtor, cctor] [dtor_ctor, ctor_dtor]));  | 
| 
49123
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents:  
diff
changeset
 | 
156  | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
157  | 
fun mk_half_distinct_tac ctxt ctor_inject abs_inject ctr_defs =  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
158  | 
  unfold_thms_tac ctxt (ctor_inject :: abs_inject :: @{thms sum.inject} @ ctr_defs) THEN
 | 
| 60728 | 159  | 
  HEADGOAL (rtac ctxt @{thm sum.distinct(1)});
 | 
| 49127 | 160  | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
161  | 
fun mk_inject_tac ctxt ctr_def ctor_inject abs_inject =  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
162  | 
unfold_thms_tac ctxt [ctr_def] THEN  | 
| 60728 | 163  | 
HEADGOAL (rtac ctxt (ctor_inject RS ssubst)) THEN  | 
| 56765 | 164  | 
  unfold_thms_tac ctxt (abs_inject :: @{thms sum.inject prod.inject conj_assoc}) THEN
 | 
| 60728 | 165  | 
HEADGOAL (rtac ctxt refl);  | 
| 49126 | 166  | 
|
| 55867 | 167  | 
val rec_unfold_thms =  | 
| 
55414
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
55083 
diff
changeset
 | 
168  | 
  @{thms comp_def convol_def fst_conv id_def case_prod_Pair_iden snd_conv split_conv
 | 
| 55966 | 169  | 
case_unit_Unity} @ sumprod_thms_map;  | 
| 49205 | 170  | 
|
| 58734 | 171  | 
fun mk_co_rec_o_map_tac ctxt co_rec_def pre_map_defs map_ident0s abs_inverses xtor_co_rec_o_map =  | 
| 58732 | 172  | 
let  | 
| 58734 | 173  | 
    val rec_o_map_simps = @{thms o_def[abs_def] id_def case_prod_app case_sum_map_sum map_sum.simps
 | 
174  | 
case_prod_map_prod id_bnf_def map_prod_simp map_sum_if_distrib_then map_sum_if_distrib_else  | 
|
175  | 
if_distrib[THEN sym]};  | 
|
| 58732 | 176  | 
in  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59270 
diff
changeset
 | 
177  | 
HEADGOAL (subst_tac ctxt (SOME [1, 2]) [co_rec_def] THEN'  | 
| 60728 | 178  | 
rtac ctxt (xtor_co_rec_o_map RS trans) THEN'  | 
| 58732 | 179  | 
CONVERSION Thm.eta_long_conversion THEN'  | 
180  | 
asm_simp_tac (ss_only (pre_map_defs @ distinct Thm.eq_thm_prop (map_ident0s @ abs_inverses) @  | 
|
181  | 
rec_o_map_simps) ctxt))  | 
|
182  | 
end;  | 
|
183  | 
||
| 64624 | 184  | 
fun mk_rec_tac pre_map_defs map_ident0s rec_defs ctor_rec pre_abs_inverse abs_inverse ctr_def ctxt =  | 
| 
58359
 
3782c7b0d1cc
avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
 
blanchet 
parents: 
58353 
diff
changeset
 | 
185  | 
HEADGOAL ((if is_def_looping ctr_def then subst_tac ctxt NONE  | 
| 
 
3782c7b0d1cc
avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
 
blanchet 
parents: 
58353 
diff
changeset
 | 
186  | 
else SELECT_GOAL o unfold_thms_tac ctxt) [ctr_def]) THEN  | 
| 64624 | 187  | 
unfold_thms_tac ctxt (ctor_rec :: pre_abs_inverse :: abs_inverse :: rec_defs @  | 
| 60728 | 188  | 
pre_map_defs @ map_ident0s @ rec_unfold_thms) THEN HEADGOAL (rtac ctxt refl);  | 
| 49205 | 189  | 
|
| 58966 | 190  | 
fun mk_rec_transfer_tac ctxt nn ns actives passives xssss rec_defs ctor_rec_transfers rel_pre_T_defs  | 
| 58446 | 191  | 
rel_eqs =  | 
192  | 
let  | 
|
193  | 
val ctor_rec_transfers' =  | 
|
| 60784 | 194  | 
map (infer_instantiate' ctxt (map SOME (passives @ actives))) ctor_rec_transfers;  | 
| 58507 | 195  | 
val total_n = Integer.sum ns;  | 
| 69593 | 196  | 
val True = \<^term>\<open>True\<close>;  | 
| 58446 | 197  | 
in  | 
198  | 
HEADGOAL Goal.conjunction_tac THEN  | 
|
199  | 
EVERY (map (fn ctor_rec_transfer =>  | 
|
| 60728 | 200  | 
REPEAT_DETERM (HEADGOAL (rtac ctxt rel_funI)) THEN  | 
| 58446 | 201  | 
unfold_thms_tac ctxt rec_defs THEN  | 
| 60728 | 202  | 
HEADGOAL (etac ctxt (mk_rel_funDN_rotated (nn + 1) ctor_rec_transfer)) THEN  | 
| 58446 | 203  | 
unfold_thms_tac ctxt rel_pre_T_defs THEN  | 
| 58966 | 204  | 
        EVERY (fst (@{fold_map 2} (fn k => fn xsss => fn acc =>
 | 
205  | 
rpair (k + acc)  | 
|
| 60728 | 206  | 
            (HEADGOAL (rtac ctxt (mk_rel_funDN_rotated 2 @{thm comp_transfer})) THEN
 | 
207  | 
             HEADGOAL (rtac ctxt @{thm vimage2p_rel_fun}) THEN
 | 
|
| 58446 | 208  | 
unfold_thms_tac ctxt rel_eqs THEN  | 
| 58966 | 209  | 
             EVERY (@{map 2} (fn n => fn xss =>
 | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59270 
diff
changeset
 | 
210  | 
REPEAT_DETERM (HEADGOAL (resolve_tac ctxt  | 
| 58966 | 211  | 
[mk_rel_funDN 2 case_sum_transfer_eq, mk_rel_funDN 2 case_sum_transfer])) THEN  | 
| 60728 | 212  | 
HEADGOAL (select_prem_tac ctxt total_n (dtac ctxt asm_rl) (acc + n)) THEN  | 
| 58446 | 213  | 
HEADGOAL (SELECT_GOAL (HEADGOAL  | 
| 60752 | 214  | 
(REPEAT_DETERM o (assume_tac ctxt ORELSE' resolve_tac ctxt  | 
| 58966 | 215  | 
[mk_rel_funDN 1 case_prod_transfer_eq,  | 
216  | 
mk_rel_funDN 1 case_prod_transfer,  | 
|
217  | 
rel_funI]) THEN_ALL_NEW  | 
|
218  | 
                    (Subgoal.FOCUS (fn {prems, ...} =>
 | 
|
219  | 
let val thm = prems  | 
|
220  | 
|> permute_like (op =) (True :: flat xss) (True :: flat_rec_arg_args xss)  | 
|
221  | 
|> Library.foldl1 (fn (acc, elem) => elem RS (acc RS rel_funD))  | 
|
| 60728 | 222  | 
in HEADGOAL (rtac ctxt thm) end) ctxt)))))  | 
| 58966 | 223  | 
(1 upto k) xsss)))  | 
224  | 
ns xssss 0)))  | 
|
| 58446 | 225  | 
ctor_rec_transfers')  | 
226  | 
end;  | 
|
227  | 
||
| 55966 | 228  | 
val corec_unfold_thms = @{thms id_def} @ sumprod_thms_map;
 | 
| 49683 | 229  | 
|
| 57399 | 230  | 
fun mk_corec_tac corec_defs map_ident0s ctor_dtor_corec pre_map_def abs_inverse ctr_def ctxt =  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
231  | 
let  | 
| 57399 | 232  | 
val ss = ss_only (pre_map_def :: abs_inverse :: map_ident0s @ corec_unfold_thms @  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
233  | 
      @{thms o_apply vimage2p_def if_True if_False}) ctxt;
 | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
234  | 
in  | 
| 55867 | 235  | 
unfold_thms_tac ctxt (ctr_def :: corec_defs) THEN  | 
| 60728 | 236  | 
HEADGOAL (rtac ctxt (ctor_dtor_corec RS trans) THEN' asm_simp_tac ss) THEN_MAYBE  | 
237  | 
    HEADGOAL (rtac ctxt refl ORELSE' rtac ctxt (@{thm unit_eq} RS arg_cong))
 | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
238  | 
end;  | 
| 49213 | 239  | 
|
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57893 
diff
changeset
 | 
240  | 
fun mk_corec_disc_iff_tac case_splits' corecs discs ctxt =  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58507 
diff
changeset
 | 
241  | 
  EVERY (@{map 3} (fn case_split_tac => fn corec_thm => fn disc =>
 | 
| 55867 | 242  | 
HEADGOAL case_split_tac THEN unfold_thms_tac ctxt [corec_thm] THEN  | 
| 52324 | 243  | 
HEADGOAL (asm_simp_tac (ss_only basic_simp_thms ctxt)) THEN  | 
| 60728 | 244  | 
(if is_refl disc then all_tac else HEADGOAL (rtac ctxt disc)))  | 
245  | 
(map (rtac ctxt) case_splits' @ [K all_tac]) corecs discs);  | 
|
| 49482 | 246  | 
|
| 58448 | 247  | 
fun mk_corec_transfer_tac ctxt actives passives type_definitions corec_defs dtor_corec_transfers  | 
248  | 
rel_pre_T_defs rel_eqs pgs pss qssss gssss =  | 
|
249  | 
let  | 
|
250  | 
val num_pgs = length pgs;  | 
|
251  | 
fun prem_no_of x = 1 + find_index (curry (op =) x) pgs;  | 
|
252  | 
||
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59270 
diff
changeset
 | 
253  | 
val Inl_Inr_Pair_tac = REPEAT_DETERM o (resolve_tac ctxt  | 
| 58968 | 254  | 
      [mk_rel_funDN 1 @{thm Inl_transfer},
 | 
| 67399 | 255  | 
       mk_rel_funDN 1 @{thm Inl_transfer[of "(=)" "(=)", simplified sum.rel_eq]},
 | 
| 58968 | 256  | 
       mk_rel_funDN 1 @{thm Inr_transfer},
 | 
| 67399 | 257  | 
       mk_rel_funDN 1 @{thm Inr_transfer[of "(=)" "(=)", simplified sum.rel_eq]},
 | 
| 58968 | 258  | 
       mk_rel_funDN 2 @{thm Pair_transfer},
 | 
| 67399 | 259  | 
       mk_rel_funDN 2 @{thm Pair_transfer[of "(=)" "(=)", simplified prod.rel_eq]}]);
 | 
| 58448 | 260  | 
|
261  | 
fun mk_unfold_If_tac total pos =  | 
|
262  | 
HEADGOAL (Inl_Inr_Pair_tac THEN'  | 
|
| 60728 | 263  | 
        rtac ctxt (mk_rel_funDN 3 @{thm If_transfer}) THEN'
 | 
264  | 
select_prem_tac ctxt total (dtac ctxt asm_rl) pos THEN'  | 
|
| 60752 | 265  | 
dtac ctxt rel_funD THEN' assume_tac ctxt THEN' assume_tac ctxt);  | 
| 58448 | 266  | 
|
267  | 
fun mk_unfold_Inl_Inr_Pair_tac total pos =  | 
|
268  | 
HEADGOAL (Inl_Inr_Pair_tac THEN'  | 
|
| 60728 | 269  | 
select_prem_tac ctxt total (dtac ctxt asm_rl) pos THEN'  | 
| 60752 | 270  | 
dtac ctxt rel_funD THEN' assume_tac ctxt THEN' assume_tac ctxt);  | 
| 58448 | 271  | 
|
272  | 
fun mk_unfold_arg_tac qs gs =  | 
|
273  | 
EVERY (map (mk_unfold_If_tac num_pgs o prem_no_of) qs) THEN  | 
|
274  | 
EVERY (map (mk_unfold_Inl_Inr_Pair_tac num_pgs o prem_no_of) gs);  | 
|
275  | 
||
276  | 
fun mk_unfold_ctr_tac type_definition qss gss =  | 
|
| 60728 | 277  | 
      HEADGOAL (rtac ctxt (mk_rel_funDN 1 (@{thm Abs_transfer} OF
 | 
| 58448 | 278  | 
[type_definition, type_definition])) THEN' Inl_Inr_Pair_tac) THEN  | 
279  | 
(case (qss, gss) of  | 
|
| 60728 | 280  | 
([], []) => HEADGOAL (rtac ctxt refl)  | 
| 58448 | 281  | 
| _ => EVERY (map2 mk_unfold_arg_tac qss gss));  | 
282  | 
||
283  | 
fun mk_unfold_type_tac type_definition ps qsss gsss =  | 
|
284  | 
let  | 
|
285  | 
val p_tacs = map (mk_unfold_If_tac num_pgs o prem_no_of) ps;  | 
|
286  | 
val qg_tacs = map2 (mk_unfold_ctr_tac type_definition) qsss gsss;  | 
|
287  | 
fun mk_unfold_ty [] [qg_tac] = qg_tac  | 
|
288  | 
| mk_unfold_ty (p_tac :: p_tacs) (qg_tac :: qg_tacs) =  | 
|
289  | 
p_tac THEN qg_tac THEN mk_unfold_ty p_tacs qg_tacs  | 
|
290  | 
in  | 
|
| 60728 | 291  | 
HEADGOAL (rtac ctxt rel_funI) THEN mk_unfold_ty p_tacs qg_tacs  | 
| 58448 | 292  | 
end;  | 
293  | 
||
294  | 
fun mk_unfold_corec_type_tac dtor_corec_transfer corec_def =  | 
|
295  | 
let  | 
|
296  | 
val active :: actives' = actives;  | 
|
| 60784 | 297  | 
val dtor_corec_transfer' =  | 
298  | 
infer_instantiate' ctxt  | 
|
299  | 
(SOME active :: map SOME passives @ map SOME actives') dtor_corec_transfer;  | 
|
| 58448 | 300  | 
in  | 
| 60728 | 301  | 
HEADGOAL Goal.conjunction_tac THEN REPEAT_DETERM (HEADGOAL (rtac ctxt rel_funI)) THEN  | 
| 58448 | 302  | 
unfold_thms_tac ctxt [corec_def] THEN  | 
| 60728 | 303  | 
HEADGOAL (etac ctxt (mk_rel_funDN_rotated (1 + length actives) dtor_corec_transfer')) THEN  | 
| 58448 | 304  | 
unfold_thms_tac ctxt (rel_pre_T_defs @ rel_eqs)  | 
305  | 
end;  | 
|
306  | 
||
307  | 
fun mk_unfold_prop_tac dtor_corec_transfer corec_def =  | 
|
308  | 
mk_unfold_corec_type_tac dtor_corec_transfer corec_def THEN  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58507 
diff
changeset
 | 
309  | 
      EVERY (@{map 4} mk_unfold_type_tac type_definitions pss qssss gssss);
 | 
| 58448 | 310  | 
in  | 
311  | 
HEADGOAL Goal.conjunction_tac THEN  | 
|
312  | 
EVERY (map2 mk_unfold_prop_tac dtor_corec_transfers corec_defs)  | 
|
313  | 
end;  | 
|
314  | 
||
| 51798 | 315  | 
fun solve_prem_prem_tac ctxt =  | 
| 61760 | 316  | 
  REPEAT o (eresolve_tac ctxt @{thms bexE rev_bexI} ORELSE'
 | 
317  | 
    rtac ctxt @{thm rev_bexI[OF UNIV_I]} ORELSE' hyp_subst_tac ctxt ORELSE'
 | 
|
318  | 
    resolve_tac ctxt @{thms disjI1 disjI2}) THEN'
 | 
|
| 60752 | 319  | 
  (rtac ctxt refl ORELSE' assume_tac ctxt ORELSE' rtac ctxt @{thm singletonI});
 | 
| 49426 | 320  | 
|
| 64624 | 321  | 
fun mk_induct_leverage_prem_prems_tac ctxt nn kks pre_abs_inverses abs_inverses set_maps  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
322  | 
pre_set_defs =  | 
| 61760 | 323  | 
HEADGOAL (EVERY' (maps (fn kk => [select_prem_tac ctxt nn (dtac ctxt meta_spec) kk,  | 
324  | 
etac ctxt meta_mp,  | 
|
| 64624 | 325  | 
SELECT_GOAL (unfold_thms_tac ctxt (pre_set_defs @ pre_abs_inverses @ abs_inverses @ set_maps @  | 
| 55966 | 326  | 
sumprod_thms_set)),  | 
| 52324 | 327  | 
solve_prem_prem_tac ctxt]) (rev kks)));  | 
| 49368 | 328  | 
|
| 64624 | 329  | 
fun mk_induct_discharge_prem_tac ctxt nn n pre_abs_inverses abs_inverses set_maps pre_set_defs m k  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
330  | 
kks =  | 
| 
49429
 
64ac3471005a
cleaner way of dealing with the set functions of sums and products
 
blanchet 
parents: 
49428 
diff
changeset
 | 
331  | 
let val r = length kks in  | 
| 60728 | 332  | 
HEADGOAL (EVERY' [select_prem_tac ctxt n (rotate_tac 1) k, rotate_tac ~1, hyp_subst_tac ctxt,  | 
333  | 
REPEAT_DETERM_N m o (dtac ctxt meta_spec THEN' rotate_tac ~1)]) THEN  | 
|
| 49391 | 334  | 
EVERY [REPEAT_DETERM_N r  | 
| 60728 | 335  | 
(HEADGOAL (rotate_tac ~1 THEN' dtac ctxt meta_mp THEN' rotate_tac 1) THEN prefer_tac 2),  | 
| 60752 | 336  | 
if r > 0 then ALLGOALS (Goal.norm_hhf_tac ctxt) else all_tac, HEADGOAL (assume_tac ctxt),  | 
| 64624 | 337  | 
mk_induct_leverage_prem_prems_tac ctxt nn kks pre_abs_inverses abs_inverses set_maps  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
338  | 
pre_set_defs]  | 
| 49391 | 339  | 
end;  | 
| 49368 | 340  | 
|
| 64624 | 341  | 
fun mk_induct_tac ctxt nn ns mss kksss ctr_defs ctor_induct' pre_abs_inverses abs_inverses set_maps  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
342  | 
pre_set_defss =  | 
| 49590 | 343  | 
let val n = Integer.sum ns in  | 
| 
58359
 
3782c7b0d1cc
avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
 
blanchet 
parents: 
58353 
diff
changeset
 | 
344  | 
(if exists is_def_looping ctr_defs then  | 
| 
 
3782c7b0d1cc
avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
 
blanchet 
parents: 
58353 
diff
changeset
 | 
345  | 
EVERY (map (fn def => HEADGOAL (subst_asm_tac ctxt NONE [def])) ctr_defs)  | 
| 
 
3782c7b0d1cc
avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
 
blanchet 
parents: 
58353 
diff
changeset
 | 
346  | 
else  | 
| 
 
3782c7b0d1cc
avoid 'subst_tac' when possible (it is suspected of not helping 'HOL-Proofs')
 
blanchet 
parents: 
58353 
diff
changeset
 | 
347  | 
unfold_thms_tac ctxt ctr_defs) THEN  | 
| 60728 | 348  | 
HEADGOAL (rtac ctxt ctor_induct') THEN co_induct_inst_as_projs_tac ctxt 0 THEN  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58507 
diff
changeset
 | 
349  | 
    EVERY (@{map 4} (EVERY oooo @{map 3} o
 | 
| 64624 | 350  | 
mk_induct_discharge_prem_tac ctxt nn n pre_abs_inverses abs_inverses set_maps)  | 
351  | 
pre_set_defss mss (unflat mss (1 upto n)) kksss)  | 
|
| 49368 | 352  | 
end;  | 
353  | 
||
| 64624 | 354  | 
fun mk_coinduct_same_ctr_tac ctxt rel_eqs pre_rel_def pre_abs_inverse abs_inverse dtor_ctor ctr_def  | 
| 64636 | 355  | 
discs sels extra_unfolds =  | 
| 51798 | 356  | 
hyp_subst_tac ctxt THEN'  | 
| 49665 | 357  | 
CONVERSION (hhf_concl_conv  | 
358  | 
(Conv.top_conv (K (Conv.try_conv (Conv.rewr_conv ctr_def))) ctxt) ctxt) THEN'  | 
|
| 
49642
 
9f884142334c
fixed simplification of prod and sum relators to avoid issues with e.g. codata ('a, 'b) k = K "'a + 'b"
 
blanchet 
parents: 
49639 
diff
changeset
 | 
359  | 
SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: dtor_ctor :: sels)) THEN'  | 
| 64624 | 360  | 
SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: pre_abs_inverse :: abs_inverse :: dtor_ctor ::  | 
| 64636 | 361  | 
    sels @ sumprod_thms_rel @ extra_unfolds @ @{thms o_apply vimage2p_def})) THEN'
 | 
| 60752 | 362  | 
(assume_tac ctxt ORELSE' REPEAT o etac ctxt conjE THEN'  | 
| 56765 | 363  | 
full_simp_tac (ss_only (no_refl discs @ rel_eqs @ more_simp_thms) ctxt) THEN'  | 
| 60728 | 364  | 
REPEAT o etac ctxt conjE THEN_MAYBE' REPEAT o hyp_subst_tac ctxt THEN'  | 
| 60752 | 365  | 
REPEAT o (resolve_tac ctxt [refl, conjI] ORELSE' assume_tac ctxt));  | 
| 49590 | 366  | 
|
| 52966 | 367  | 
fun mk_coinduct_distinct_ctrs_tac ctxt discs discs' =  | 
| 
54198
 
4fadf746f2d5
got rid of annoying duplicate rewrite rule warnings
 
blanchet 
parents: 
53690 
diff
changeset
 | 
368  | 
let  | 
| 
 
4fadf746f2d5
got rid of annoying duplicate rewrite rule warnings
 
blanchet 
parents: 
53690 
diff
changeset
 | 
369  | 
    val discs'' = map (perhaps (try (fn th => th RS @{thm notnotD}))) (discs @ discs')
 | 
| 
 
4fadf746f2d5
got rid of annoying duplicate rewrite rule warnings
 
blanchet 
parents: 
53690 
diff
changeset
 | 
370  | 
|> distinct Thm.eq_thm_prop;  | 
| 
 
4fadf746f2d5
got rid of annoying duplicate rewrite rule warnings
 
blanchet 
parents: 
53690 
diff
changeset
 | 
371  | 
in  | 
| 60728 | 372  | 
hyp_subst_tac ctxt THEN' REPEAT o etac ctxt conjE THEN'  | 
| 
54198
 
4fadf746f2d5
got rid of annoying duplicate rewrite rule warnings
 
blanchet 
parents: 
53690 
diff
changeset
 | 
373  | 
full_simp_tac (ss_only (refl :: no_refl discs'' @ basic_simp_thms) ctxt)  | 
| 
 
4fadf746f2d5
got rid of annoying duplicate rewrite rule warnings
 
blanchet 
parents: 
53690 
diff
changeset
 | 
374  | 
end;  | 
| 49590 | 375  | 
|
| 64636 | 376  | 
fun mk_coinduct_discharge_prem_tac ctxt extra_unfolds rel_eqs' nn kk n pre_rel_def pre_abs_inverse  | 
377  | 
abs_inverse dtor_ctor exhaust ctr_defs discss selss =  | 
|
| 49590 | 378  | 
let val ks = 1 upto n in  | 
| 61760 | 379  | 
EVERY' ([rtac ctxt allI, rtac ctxt allI, rtac ctxt impI,  | 
380  | 
select_prem_tac ctxt nn (dtac ctxt meta_spec) kk, dtac ctxt meta_spec, dtac ctxt meta_mp,  | 
|
381  | 
assume_tac ctxt, rtac ctxt exhaust, K (co_induct_inst_as_projs_tac ctxt 0),  | 
|
382  | 
hyp_subst_tac ctxt] @  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58507 
diff
changeset
 | 
383  | 
      @{map 4} (fn k => fn ctr_def => fn discs => fn sels =>
 | 
| 60728 | 384  | 
EVERY' ([rtac ctxt exhaust, K (co_induct_inst_as_projs_tac ctxt 1)] @  | 
| 
49591
 
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
 
blanchet 
parents: 
49590 
diff
changeset
 | 
385  | 
map2 (fn k' => fn discs' =>  | 
| 
 
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
 
blanchet 
parents: 
49590 
diff
changeset
 | 
386  | 
if k' = k then  | 
| 64624 | 387  | 
mk_coinduct_same_ctr_tac ctxt rel_eqs' pre_rel_def pre_abs_inverse abs_inverse  | 
| 64636 | 388  | 
dtor_ctor ctr_def discs sels extra_unfolds  | 
| 
49591
 
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
 
blanchet 
parents: 
49590 
diff
changeset
 | 
389  | 
else  | 
| 52966 | 390  | 
mk_coinduct_distinct_ctrs_tac ctxt discs discs') ks discss)) ks ctr_defs discss selss)  | 
| 49590 | 391  | 
end;  | 
392  | 
||
| 64624 | 393  | 
fun mk_coinduct_tac ctxt rel_eqs' nn ns dtor_coinduct' pre_rel_defs pre_abs_inverses abs_inverses  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55642 
diff
changeset
 | 
394  | 
dtor_ctors exhausts ctr_defss discsss selsss =  | 
| 60728 | 395  | 
HEADGOAL (rtac ctxt dtor_coinduct' THEN'  | 
| 64636 | 396  | 
    EVERY' (@{map 10} (mk_coinduct_discharge_prem_tac ctxt [] rel_eqs' nn)
 | 
| 64624 | 397  | 
(1 upto nn) ns pre_rel_defs pre_abs_inverses abs_inverses dtor_ctors exhausts ctr_defss  | 
398  | 
discsss selsss));  | 
|
| 49590 | 399  | 
|
| 
64415
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64067 
diff
changeset
 | 
400  | 
fun mk_map_tac ctxt abs_inverses pre_map_def map_ctor live_nesting_map_id0s ctr_defs'  | 
| 
63845
 
61a03e429cbd
generalized code towards nonuniform (co)datatypes
 
blanchet 
parents: 
63842 
diff
changeset
 | 
401  | 
extra_unfolds =  | 
| 63840 | 402  | 
TRYALL Goal.conjunction_tac THEN  | 
| 
64415
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64067 
diff
changeset
 | 
403  | 
unfold_thms_tac ctxt (pre_map_def :: map_ctor :: abs_inverses @ live_nesting_map_id0s @  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64067 
diff
changeset
 | 
404  | 
ctr_defs' @ extra_unfolds @ sumprod_thms_map @  | 
| 64067 | 405  | 
    @{thms o_apply id_apply id_o o_id}) THEN
 | 
| 63840 | 406  | 
ALLGOALS (rtac ctxt refl);  | 
407  | 
||
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57893 
diff
changeset
 | 
408  | 
fun mk_map_disc_iff_tac ctxt ct exhaust discs maps =  | 
| 
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57893 
diff
changeset
 | 
409  | 
TRYALL Goal.conjunction_tac THEN  | 
| 60784 | 410  | 
ALLGOALS (rtac ctxt (infer_instantiate' ctxt [SOME ct] exhaust) THEN_ALL_NEW  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57893 
diff
changeset
 | 
411  | 
REPEAT_DETERM o hyp_subst_tac ctxt) THEN  | 
| 
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57893 
diff
changeset
 | 
412  | 
unfold_thms_tac ctxt maps THEN  | 
| 
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57893 
diff
changeset
 | 
413  | 
unfold_thms_tac ctxt (map (fn thm => thm RS eqFalseI  | 
| 
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57893 
diff
changeset
 | 
414  | 
handle THM _ => thm RS eqTrueI) discs) THEN  | 
| 60728 | 415  | 
ALLGOALS (rtac ctxt refl ORELSE' rtac ctxt TrueI);  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57893 
diff
changeset
 | 
416  | 
|
| 
58326
 
7e142efcee1a
make 'rel_sel' and 'map_sel' tactics more robust
 
desharna 
parents: 
58181 
diff
changeset
 | 
417  | 
fun mk_map_sel_tac ctxt ct exhaust discs maps sels map_id0s =  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57893 
diff
changeset
 | 
418  | 
TRYALL Goal.conjunction_tac THEN  | 
| 63841 | 419  | 
ALLGOALS (rtac ctxt (infer_instantiate' ctxt [SOME ct] exhaust) THEN_ALL_NEW  | 
420  | 
REPEAT_DETERM o hyp_subst_tac ctxt) THEN  | 
|
421  | 
unfold_thms_tac ctxt ((discs RL [eqTrueI, eqFalseI]) @  | 
|
422  | 
    @{thms not_True_eq_False not_False_eq_True}) THEN
 | 
|
423  | 
  TRYALL (etac ctxt FalseE ORELSE' etac ctxt @{thm TrueE}) THEN
 | 
|
424  | 
  unfold_thms_tac ctxt (@{thm id_apply} :: maps @ sels @ map_id0s) THEN
 | 
|
425  | 
ALLGOALS (rtac ctxt refl);  | 
|
426  | 
||
| 
64415
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64067 
diff
changeset
 | 
427  | 
fun mk_rel_tac ctxt abs_inverses pre_rel_def rel_ctor live_nesting_rel_eqs ctr_defs' extra_unfolds =  | 
| 63841 | 428  | 
TRYALL Goal.conjunction_tac THEN  | 
| 
64415
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64067 
diff
changeset
 | 
429  | 
unfold_thms_tac ctxt (pre_rel_def :: rel_ctor :: abs_inverses @ live_nesting_rel_eqs @ ctr_defs' @  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64067 
diff
changeset
 | 
430  | 
    extra_unfolds @ sumprod_thms_rel @ @{thms vimage2p_def o_apply sum.inject
 | 
| 63842 | 431  | 
sum.distinct(1)[THEN eq_False[THEN iffD2]] not_False_eq_True}) THEN  | 
| 63841 | 432  | 
ALLGOALS (resolve_tac ctxt [TrueI, refl]);  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57893 
diff
changeset
 | 
433  | 
|
| 62335 | 434  | 
fun mk_rel_case_tac ctxt ct1 ct2 exhaust injects rel_injects distincts rel_distincts rel_eqs =  | 
| 60784 | 435  | 
HEADGOAL (rtac ctxt (infer_instantiate' ctxt [SOME ct1] exhaust) THEN_ALL_NEW  | 
436  | 
rtac ctxt (infer_instantiate' ctxt [SOME ct2] exhaust) THEN_ALL_NEW  | 
|
| 57525 | 437  | 
hyp_subst_tac ctxt) THEN  | 
| 61344 | 438  | 
unfold_thms_tac ctxt (rel_eqs @ injects @ rel_injects @  | 
439  | 
    @{thms conj_imp_eq_imp_imp simp_thms(6) True_implies_equals} @
 | 
|
| 57529 | 440  | 
map (fn thm => thm RS eqFalseI) (distincts @ rel_distincts) @  | 
441  | 
map (fn thm => thm RS eqTrueI) rel_injects) THEN  | 
|
| 60752 | 442  | 
TRYALL (assume_tac ctxt ORELSE' etac ctxt FalseE ORELSE'  | 
| 61344 | 443  | 
(REPEAT_DETERM o dtac ctxt meta_spec THEN'  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59270 
diff
changeset
 | 
444  | 
TRY o filter_prems_tac ctxt  | 
| 69593 | 445  | 
(forall (curry (op <>) (HOLogic.mk_Trueprop \<^term>\<open>False\<close>)) o Logic.strip_imp_prems) THEN'  | 
| 61344 | 446  | 
REPEAT_DETERM o (dtac ctxt meta_mp THEN' rtac ctxt refl) THEN'  | 
447  | 
(assume_tac ctxt ORELSE' Goal.assume_rule_tac ctxt)));  | 
|
| 57525 | 448  | 
|
| 
57301
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57152 
diff
changeset
 | 
449  | 
fun mk_rel_coinduct0_tac ctxt dtor_rel_coinduct cts assms exhausts discss selss ctor_defss  | 
| 
57558
 
6bb3dd7f8097
took out 'rel_cases' for now because of failing tactic
 
blanchet 
parents: 
57529 
diff
changeset
 | 
450  | 
dtor_ctors ctor_injects abs_injects rel_pre_defs abs_inverses nesting_rel_eqs =  | 
| 60728 | 451  | 
rtac ctxt dtor_rel_coinduct 1 THEN  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58507 
diff
changeset
 | 
452  | 
   EVERY (@{map 11} (fn ct => fn assm => fn exhaust => fn discs => fn sels => fn ctor_defs =>
 | 
| 57670 | 453  | 
fn dtor_ctor => fn ctor_inject => fn abs_inject => fn rel_pre_def => fn abs_inverse =>  | 
| 60728 | 454  | 
(rtac ctxt exhaust THEN_ALL_NEW (rtac ctxt exhaust THEN_ALL_NEW  | 
| 60784 | 455  | 
(dtac ctxt (rotate_prems ~1 (infer_instantiate' ctxt [NONE, NONE, NONE, NONE, SOME ct]  | 
| 57670 | 456  | 
            @{thm arg_cong2} RS iffD1)) THEN'
 | 
| 60752 | 457  | 
assume_tac ctxt THEN' assume_tac ctxt THEN' hyp_subst_tac ctxt THEN' dtac ctxt assm THEN'  | 
| 60728 | 458  | 
REPEAT_DETERM o etac ctxt conjE))) 1 THEN  | 
| 57668 | 459  | 
unfold_thms_tac ctxt ((discs RL [eqTrueI, eqFalseI]) @ sels @ simp_thms') THEN  | 
| 57528 | 460  | 
unfold_thms_tac ctxt (dtor_ctor :: rel_pre_def :: abs_inverse :: ctor_inject ::  | 
| 58128 | 461  | 
abs_inject :: ctor_defs @ nesting_rel_eqs @ simp_thms' @  | 
| 62335 | 462  | 
        @{thms id_bnf_def rel_sum_simps rel_prod_inject vimage2p_def Inl_Inr_False
 | 
| 58353 | 463  | 
iffD2[OF eq_False Inr_not_Inl] sum.inject prod.inject}) THEN  | 
| 61760 | 464  | 
REPEAT_DETERM (HEADGOAL ((REPEAT_DETERM o etac ctxt conjE) THEN'  | 
465  | 
(REPEAT_DETERM o rtac ctxt conjI) THEN' (rtac ctxt refl ORELSE' assume_tac ctxt))))  | 
|
| 
57301
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57152 
diff
changeset
 | 
466  | 
cts assms exhausts discss selss ctor_defss dtor_ctors ctor_injects abs_injects rel_pre_defs  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57152 
diff
changeset
 | 
467  | 
abs_inverses);  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57152 
diff
changeset
 | 
468  | 
|
| 57471 | 469  | 
fun mk_rel_induct0_tac ctxt ctor_rel_induct assms cterms exhausts ctor_defss ctor_injects  | 
470  | 
rel_pre_list_defs Abs_inverses nesting_rel_eqs =  | 
|
| 60728 | 471  | 
  rtac ctxt ctor_rel_induct 1 THEN EVERY (@{map 6} (fn cterm => fn exhaust => fn ctor_defs =>
 | 
| 57471 | 472  | 
fn ctor_inject => fn rel_pre_list_def => fn Abs_inverse =>  | 
| 60728 | 473  | 
HEADGOAL (rtac ctxt exhaust THEN_ALL_NEW (rtac ctxt exhaust THEN_ALL_NEW  | 
| 61760 | 474  | 
          (rtac ctxt (infer_instantiate' ctxt (replicate 4 NONE @ [SOME cterm]) @{thm arg_cong2}
 | 
475  | 
RS iffD2)  | 
|
| 60752 | 476  | 
THEN' assume_tac ctxt THEN' assume_tac ctxt THEN' TRY o resolve_tac ctxt assms))) THEN  | 
| 57700 | 477  | 
unfold_thms_tac ctxt (ctor_inject :: rel_pre_list_def :: ctor_defs @ nesting_rel_eqs @  | 
| 58353 | 478  | 
          @{thms id_bnf_def vimage2p_def}) THEN
 | 
| 57471 | 479  | 
TRYALL (hyp_subst_tac ctxt) THEN  | 
| 62335 | 480  | 
        unfold_thms_tac ctxt (Abs_inverse :: @{thms rel_sum_simps rel_prod_inject Inl_Inr_False
 | 
| 57471 | 481  | 
Inr_Inl_False sum.inject prod.inject}) THEN  | 
| 60752 | 482  | 
TRYALL (rtac ctxt refl ORELSE' etac ctxt FalseE ORELSE'  | 
483  | 
(REPEAT_DETERM o etac ctxt conjE) THEN' assume_tac ctxt))  | 
|
| 57471 | 484  | 
cterms exhausts ctor_defss ctor_injects rel_pre_list_defs Abs_inverses);  | 
485  | 
||
| 
58326
 
7e142efcee1a
make 'rel_sel' and 'map_sel' tactics more robust
 
desharna 
parents: 
58181 
diff
changeset
 | 
486  | 
fun mk_rel_sel_tac ctxt ct1 ct2 exhaust discs sels rel_injects distincts rel_distincts rel_eqs =  | 
| 60784 | 487  | 
HEADGOAL (rtac ctxt (infer_instantiate' ctxt [SOME ct1] exhaust) THEN_ALL_NEW  | 
| 63851 | 488  | 
rtac ctxt (infer_instantiate' ctxt [SOME ct2] exhaust) THEN_ALL_NEW hyp_subst_tac ctxt) THEN  | 
| 59270 | 489  | 
unfold_thms_tac ctxt (sels @ rel_injects @ rel_eqs @  | 
490  | 
    @{thms simp_thms(6,7,8,11,12,15,16,21,22,24)} @ ((discs @ distincts) RL [eqTrueI, eqFalseI]) @
 | 
|
491  | 
(rel_injects RL [eqTrueI]) @ (rel_distincts RL [eqFalseI])) THEN  | 
|
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59270 
diff
changeset
 | 
492  | 
TRYALL (resolve_tac ctxt [TrueI, refl]);  | 
| 57563 | 493  | 
|
| 58676 | 494  | 
fun mk_sel_transfer_tac ctxt n sel_defs case_transfer =  | 
495  | 
TRYALL Goal.conjunction_tac THEN  | 
|
| 
63068
 
8b9401bfd9fd
unfold is subject to unfold_abs_def (still inactive);
 
wenzelm 
parents: 
62535 
diff
changeset
 | 
496  | 
unfold_thms_tac ctxt (map (Local_Defs.abs_def_rule ctxt) sel_defs) THEN  | 
| 60728 | 497  | 
ALLGOALS (rtac ctxt (mk_rel_funDN n case_transfer) THEN_ALL_NEW  | 
| 60752 | 498  | 
REPEAT_DETERM o (assume_tac ctxt ORELSE' rtac ctxt rel_funI));  | 
| 58676 | 499  | 
|
| 63851 | 500  | 
fun mk_set0_tac ctxt abs_inverses pre_set_defs dtor_ctor fp_sets fp_nesting_set_maps  | 
501  | 
live_nesting_set_maps ctr_defs' extra_unfolds =  | 
|
502  | 
TRYALL Goal.conjunction_tac THEN  | 
|
| 63852 | 503  | 
unfold_thms_tac ctxt ctr_defs' THEN  | 
504  | 
ALLGOALS (subst_tac ctxt NONE fp_sets) THEN  | 
|
505  | 
unfold_thms_tac ctxt (dtor_ctor :: abs_inverses @ pre_set_defs @ fp_nesting_set_maps @  | 
|
506  | 
live_nesting_set_maps @ extra_unfolds @ basic_sumprod_thms_set @  | 
|
| 63854 | 507  | 
    @{thms UN_UN_flatten UN_Un_distrib UN_Un sup_assoc[THEN sym]}) THEN
 | 
| 63852 | 508  | 
  ALLGOALS (rtac ctxt @{thm set_eqI[OF iffI]}) THEN
 | 
509  | 
ALLGOALS (REPEAT_DETERM o etac ctxt UnE) THEN  | 
|
510  | 
  ALLGOALS (REPEAT o resolve_tac ctxt @{thms UnI1 UnI2} THEN' assume_tac ctxt);
 | 
|
| 63851 | 511  | 
|
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57893 
diff
changeset
 | 
512  | 
fun mk_set_sel_tac ctxt ct exhaust discs sels sets =  | 
| 57152 | 513  | 
TRYALL Goal.conjunction_tac THEN  | 
| 63841 | 514  | 
ALLGOALS (rtac ctxt (infer_instantiate' ctxt [SOME ct] exhaust) THEN_ALL_NEW  | 
515  | 
REPEAT_DETERM o hyp_subst_tac ctxt) THEN  | 
|
516  | 
unfold_thms_tac ctxt ((discs RL [eqTrueI, eqFalseI]) @  | 
|
517  | 
    @{thms not_True_eq_False not_False_eq_True}) THEN
 | 
|
518  | 
  TRYALL (etac ctxt FalseE ORELSE' etac ctxt @{thm TrueE}) THEN
 | 
|
519  | 
unfold_thms_tac ctxt (sels @ sets) THEN  | 
|
520  | 
  ALLGOALS (REPEAT o (resolve_tac ctxt @{thms UnI1 UnI2 imageI} ORELSE'
 | 
|
521  | 
      eresolve_tac ctxt @{thms UN_I UN_I[rotated] imageE} ORELSE'
 | 
|
522  | 
hyp_subst_tac ctxt) THEN'  | 
|
523  | 
    (rtac ctxt @{thm singletonI} ORELSE' assume_tac ctxt));
 | 
|
| 57152 | 524  | 
|
| 57893 | 525  | 
fun mk_set_cases_tac ctxt ct assms exhaust sets =  | 
| 61760 | 526  | 
HEADGOAL (rtac ctxt (infer_instantiate' ctxt [SOME ct] exhaust)  | 
527  | 
THEN_ALL_NEW hyp_subst_tac ctxt) THEN  | 
|
| 57893 | 528  | 
unfold_thms_tac ctxt sets THEN  | 
529  | 
REPEAT_DETERM (HEADGOAL  | 
|
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59270 
diff
changeset
 | 
530  | 
    (eresolve_tac ctxt @{thms FalseE emptyE singletonE UnE UN_E insertE} ORELSE'
 | 
| 57893 | 531  | 
hyp_subst_tac ctxt ORELSE'  | 
| 61760 | 532  | 
SELECT_GOAL (SOLVE (HEADGOAL (eresolve_tac ctxt assms THEN' REPEAT_DETERM o  | 
533  | 
assume_tac ctxt)))));  | 
|
| 57893 | 534  | 
|
| 57891 | 535  | 
fun mk_set_intros_tac ctxt sets =  | 
| 
59856
 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 
blanchet 
parents: 
59794 
diff
changeset
 | 
536  | 
TRYALL Goal.conjunction_tac THEN unfold_thms_tac ctxt sets THEN  | 
| 57891 | 537  | 
TRYALL (REPEAT o  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59270 
diff
changeset
 | 
538  | 
    (resolve_tac ctxt @{thms UnI1 UnI2} ORELSE'
 | 
| 60752 | 539  | 
     eresolve_tac ctxt @{thms UN_I UN_I[rotated]}) THEN'
 | 
540  | 
     (rtac ctxt @{thm singletonI} ORELSE' assume_tac ctxt));
 | 
|
| 57891 | 541  | 
|
| 57700 | 542  | 
fun mk_set_induct0_tac ctxt cts assms dtor_set_inducts exhausts set_pre_defs ctor_defs dtor_ctors  | 
543  | 
Abs_pre_inverses =  | 
|
544  | 
let  | 
|
| 
58417
 
fa50722ad6cb
make 'set_induct0' tactic more robust w.r.t multiple arguments constructors
 
desharna 
parents: 
58359 
diff
changeset
 | 
545  | 
val assms_tac =  | 
| 
 
fa50722ad6cb
make 'set_induct0' tactic more robust w.r.t multiple arguments constructors
 
desharna 
parents: 
58359 
diff
changeset
 | 
546  | 
      let val assms' = map (unfold_thms ctxt (@{thm id_bnf_def} :: ctor_defs)) assms in
 | 
| 
 
fa50722ad6cb
make 'set_induct0' tactic more robust w.r.t multiple arguments constructors
 
desharna 
parents: 
58359 
diff
changeset
 | 
547  | 
fold (curry (op ORELSE')) (map (fn thm =>  | 
| 61760 | 548  | 
funpow (length (Thm.prems_of thm)) (fn tac => tac THEN' assume_tac ctxt)  | 
549  | 
(rtac ctxt thm)) assms')  | 
|
| 60728 | 550  | 
(etac ctxt FalseE)  | 
| 
58417
 
fa50722ad6cb
make 'set_induct0' tactic more robust w.r.t multiple arguments constructors
 
desharna 
parents: 
58359 
diff
changeset
 | 
551  | 
end;  | 
| 57700 | 552  | 
    val exhausts' = map (fn thm => thm RS @{thm asm_rl[of "P x y" for P x y]}) exhausts
 | 
| 60784 | 553  | 
|> map2 (fn ct => infer_instantiate' ctxt [NONE, SOME ct]) cts;  | 
| 57700 | 554  | 
in  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59270 
diff
changeset
 | 
555  | 
ALLGOALS (resolve_tac ctxt dtor_set_inducts) THEN  | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59270 
diff
changeset
 | 
556  | 
TRYALL (resolve_tac ctxt exhausts' THEN_ALL_NEW  | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59270 
diff
changeset
 | 
557  | 
(resolve_tac ctxt (map (fn ct => refl RS  | 
| 60784 | 558  | 
         infer_instantiate' ctxt (replicate 4 NONE @ [SOME ct]) @{thm arg_cong2} RS iffD2) cts)
 | 
| 60752 | 559  | 
THEN' assume_tac ctxt THEN' hyp_subst_tac ctxt)) THEN  | 
| 57700 | 560  | 
unfold_thms_tac ctxt (Abs_pre_inverses @ dtor_ctors @ set_pre_defs @ ctor_defs @  | 
| 58353 | 561  | 
      @{thms id_bnf_def o_apply sum_set_simps prod_set_simps UN_empty UN_insert Un_empty_left
 | 
562  | 
Un_empty_right empty_iff singleton_iff}) THEN  | 
|
| 61760 | 563  | 
REPEAT (HEADGOAL (hyp_subst_tac ctxt ORELSE'  | 
564  | 
      eresolve_tac ctxt @{thms UN_E UnE singletonE} ORELSE' assms_tac))
 | 
|
| 57700 | 565  | 
end;  | 
566  | 
||
| 
49123
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents:  
diff
changeset
 | 
567  | 
end;  |