author | desharna |
Wed, 30 Jul 2014 10:50:28 +0200 | |
changeset 57700 | a2c4adb839a9 |
parent 57670 | d7b15b99f93c |
child 57815 | f97643a56615 |
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 |
55966 | 11 |
val sumprod_thms_map: thm list |
12 |
val sumprod_thms_set: thm list |
|
13 |
val sumprod_thms_rel: thm list |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49542
diff
changeset
|
14 |
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49590
diff
changeset
|
15 |
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
|
16 |
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
|
17 |
thm list list list -> tactic |
55867 | 18 |
val mk_corec_tac: thm list -> thm list -> thm -> thm -> thm -> thm -> Proof.context -> tactic |
49501 | 19 |
val mk_ctor_iff_dtor_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm -> |
20 |
tactic |
|
55867 | 21 |
val mk_disc_corec_iff_tac: thm list -> thm list -> thm list -> Proof.context -> tactic |
56991
8e9ca31e9b8e
generate 'disc_map_iff[simp]' theorem for (co)datatypes
desharna
parents:
56990
diff
changeset
|
22 |
val mk_disc_map_iff_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> tactic |
49161
a8e74375d971
fixed (n + 1)st bug in "mk_exhaust_tac" -- arose with uncurried constructors
blanchet
parents:
49160
diff
changeset
|
23 |
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
|
24 |
val mk_half_distinct_tac: Proof.context -> thm -> thm -> thm list -> tactic |
49590 | 25 |
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
|
26 |
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
|
27 |
val mk_inject_tac: Proof.context -> thm -> thm -> thm -> tactic |
55867 | 28 |
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
|
29 |
tactic |
57525 | 30 |
val mk_rel_cases_tac: Proof.context -> cterm -> cterm -> thm -> thm list -> thm list -> |
31 |
thm list -> thm list -> tactic |
|
57301
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57152
diff
changeset
|
32 |
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
|
33 |
thm list list -> thm list list -> thm list list -> thm list -> thm list -> thm list -> |
57471 | 34 |
thm list -> thm list -> thm list -> tactic |
35 |
val mk_rel_induct0_tac: Proof.context -> thm -> thm list -> cterm list -> thm list -> |
|
36 |
thm list list -> thm list -> thm list -> thm list -> thm list -> tactic |
|
57563 | 37 |
val mk_rel_sel_tac: Proof.context -> cterm -> cterm -> thm -> thm list -> thm list -> thm list -> |
38 |
thm list -> thm list -> tactic |
|
57046
b3613d7e108b
generate 'sel_map[simp]' theorem for (co)datatypes and tuning 'disc_map_iff'
desharna
parents:
56991
diff
changeset
|
39 |
val mk_sel_map_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> thm list -> tactic |
57152 | 40 |
val mk_sel_set_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> thm list -> tactic |
56956 | 41 |
val mk_set_empty_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> tactic |
57700 | 42 |
val mk_set_induct0_tac: Proof.context -> cterm list -> thm list -> thm list -> thm list -> |
43 |
thm list -> thm list -> thm list -> thm list -> tactic |
|
49123
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff
changeset
|
44 |
end; |
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff
changeset
|
45 |
|
49636 | 46 |
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
|
47 |
struct |
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff
changeset
|
48 |
|
56991
8e9ca31e9b8e
generate 'disc_map_iff[simp]' theorem for (co)datatypes
desharna
parents:
56990
diff
changeset
|
49 |
open Ctr_Sugar_Util |
49125 | 50 |
open BNF_Tactics |
49123
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff
changeset
|
51 |
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
|
52 |
open BNF_FP_Util |
49123
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff
changeset
|
53 |
|
49590 | 54 |
val basic_simp_thms = @{thms simp_thms(7,8,12,14,22,24)}; |
55 |
val more_simp_thms = basic_simp_thms @ @{thms simp_thms(11,15,16,21)}; |
|
57303 | 56 |
val simp_thms' = @{thms simp_thms(6,7,8,11,12,15,16,22,24)}; |
49590 | 57 |
|
55966 | 58 |
val sumprod_thms_map = @{thms id_apply map_prod_simp prod.case sum.case map_sum.simps}; |
59 |
val sumprod_thms_set = |
|
55930
25a90cebbbe5
more careful simplification of sets (cf. abf91ebd0820)---yields smaller terms
traytel
parents:
55906
diff
changeset
|
60 |
@{thms UN_empty UN_insert Un_empty_left Un_empty_right Un_iff UN_simps(10) UN_iff |
55932 | 61 |
Union_Un_distrib image_iff o_apply map_prod_simp |
55931 | 62 |
mem_Collect_eq prod_set_simps map_sum.simps sum_set_simps}; |
56765 | 63 |
val sumprod_thms_rel = @{thms rel_sum_simps rel_prod_apply prod.inject id_apply conj_assoc}; |
49368 | 64 |
|
49668 | 65 |
fun hhf_concl_conv cv ctxt ct = |
66 |
(case Thm.term_of ct of |
|
56245 | 67 |
Const (@{const_name Pure.all}, _) $ Abs _ => |
49668 | 68 |
Conv.arg_conv (Conv.abs_conv (hhf_concl_conv cv o snd) ctxt) ct |
69 |
| _ => Conv.concl_conv ~1 cv ct); |
|
70 |
||
54922 | 71 |
fun co_induct_inst_as_projs ctxt k thm = |
49368 | 72 |
let |
54922 | 73 |
val fs = Term.add_vars (prop_of thm) [] |
49368 | 74 |
|> filter (fn (_, Type (@{type_name fun}, [_, T'])) => T' <> HOLogic.boolT | _ => false); |
54923
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
blanchet
parents:
54922
diff
changeset
|
75 |
fun mk_cfp (f as (_, T)) = |
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
blanchet
parents:
54922
diff
changeset
|
76 |
(certify ctxt (Var f), certify ctxt (mk_proj T (num_binder_types T) k)); |
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
blanchet
parents:
54922
diff
changeset
|
77 |
val cfps = map mk_cfp fs; |
49368 | 78 |
in |
54923
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
blanchet
parents:
54922
diff
changeset
|
79 |
Drule.cterm_instantiate cfps thm |
49368 | 80 |
end; |
81 |
||
54922 | 82 |
val co_induct_inst_as_projs_tac = PRIMITIVE oo co_induct_inst_as_projs; |
49368 | 83 |
|
49501 | 84 |
fun mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor sumEN' = |
52324 | 85 |
unfold_thms_tac ctxt (ctor_iff_dtor :: ctr_defs) THEN HEADGOAL (rtac sumEN') THEN |
86 |
HEADGOAL (EVERY' (maps (fn k => [select_prem_tac n (rotate_tac 1) k, |
|
87 |
REPEAT_DETERM o dtac meta_spec, etac meta_mp, atac]) (1 upto n))); |
|
49125 | 88 |
|
49501 | 89 |
fun mk_ctor_iff_dtor_tac ctxt cTs cctor cdtor ctor_dtor dtor_ctor = |
52324 | 90 |
HEADGOAL (rtac iffI THEN' |
91 |
EVERY' (map3 (fn cTs => fn cx => fn th => |
|
92 |
dtac (Drule.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN' |
|
93 |
SELECT_GOAL (unfold_thms_tac ctxt [th]) THEN' |
|
94 |
atac) [rev cTs, cTs] [cdtor, cctor] [dtor_ctor, ctor_dtor])); |
|
49123
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff
changeset
|
95 |
|
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
96 |
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
|
97 |
unfold_thms_tac ctxt (ctor_inject :: abs_inject :: @{thms sum.inject} @ ctr_defs) THEN |
52324 | 98 |
HEADGOAL (rtac @{thm sum.distinct(1)}); |
49127 | 99 |
|
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
100 |
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
|
101 |
unfold_thms_tac ctxt [ctr_def] THEN |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
102 |
HEADGOAL (rtac (ctor_inject RS ssubst)) THEN |
56765 | 103 |
unfold_thms_tac ctxt (abs_inject :: @{thms sum.inject prod.inject conj_assoc}) THEN |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
104 |
HEADGOAL (rtac refl); |
49126 | 105 |
|
55867 | 106 |
val rec_unfold_thms = |
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
55083
diff
changeset
|
107 |
@{thms comp_def convol_def fst_conv id_def case_prod_Pair_iden snd_conv split_conv |
55966 | 108 |
case_unit_Unity} @ sumprod_thms_map; |
49205 | 109 |
|
57399 | 110 |
fun mk_rec_tac pre_map_defs map_ident0s rec_defs ctor_rec fp_abs_inverse abs_inverse ctr_def ctxt = |
55867 | 111 |
unfold_thms_tac ctxt (ctr_def :: ctor_rec :: fp_abs_inverse :: abs_inverse :: rec_defs @ |
57399 | 112 |
pre_map_defs @ map_ident0s @ rec_unfold_thms) THEN HEADGOAL (rtac refl); |
49205 | 113 |
|
55966 | 114 |
val corec_unfold_thms = @{thms id_def} @ sumprod_thms_map; |
49683 | 115 |
|
57399 | 116 |
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
|
117 |
let |
57399 | 118 |
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
|
119 |
@{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
|
120 |
in |
55867 | 121 |
unfold_thms_tac ctxt (ctr_def :: corec_defs) THEN |
122 |
HEADGOAL (rtac (ctor_dtor_corec RS trans) THEN' asm_simp_tac ss) THEN_MAYBE |
|
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
123 |
HEADGOAL (rtac refl ORELSE' rtac (@{thm unit_eq} RS arg_cong)) |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
124 |
end; |
49213 | 125 |
|
55867 | 126 |
fun mk_disc_corec_iff_tac case_splits' corecs discs ctxt = |
127 |
EVERY (map3 (fn case_split_tac => fn corec_thm => fn disc => |
|
128 |
HEADGOAL case_split_tac THEN unfold_thms_tac ctxt [corec_thm] THEN |
|
52324 | 129 |
HEADGOAL (asm_simp_tac (ss_only basic_simp_thms ctxt)) THEN |
130 |
(if is_refl disc then all_tac else HEADGOAL (rtac disc))) |
|
55867 | 131 |
(map rtac case_splits' @ [K all_tac]) corecs discs); |
49482 | 132 |
|
56991
8e9ca31e9b8e
generate 'disc_map_iff[simp]' theorem for (co)datatypes
desharna
parents:
56990
diff
changeset
|
133 |
fun mk_disc_map_iff_tac ctxt ct exhaust discs maps = |
8e9ca31e9b8e
generate 'disc_map_iff[simp]' theorem for (co)datatypes
desharna
parents:
56990
diff
changeset
|
134 |
TRYALL Goal.conjunction_tac THEN |
57558
6bb3dd7f8097
took out 'rel_cases' for now because of failing tactic
blanchet
parents:
57529
diff
changeset
|
135 |
ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW |
6bb3dd7f8097
took out 'rel_cases' for now because of failing tactic
blanchet
parents:
57529
diff
changeset
|
136 |
REPEAT_DETERM o hyp_subst_tac ctxt) THEN |
6bb3dd7f8097
took out 'rel_cases' for now because of failing tactic
blanchet
parents:
57529
diff
changeset
|
137 |
unfold_thms_tac ctxt maps THEN |
6bb3dd7f8097
took out 'rel_cases' for now because of failing tactic
blanchet
parents:
57529
diff
changeset
|
138 |
unfold_thms_tac ctxt (map (fn thm => thm RS eqFalseI |
6bb3dd7f8097
took out 'rel_cases' for now because of failing tactic
blanchet
parents:
57529
diff
changeset
|
139 |
handle THM _ => thm RS eqTrueI) discs) THEN |
6bb3dd7f8097
took out 'rel_cases' for now because of failing tactic
blanchet
parents:
57529
diff
changeset
|
140 |
ALLGOALS (rtac refl ORELSE' rtac TrueI); |
56991
8e9ca31e9b8e
generate 'disc_map_iff[simp]' theorem for (co)datatypes
desharna
parents:
56990
diff
changeset
|
141 |
|
51798 | 142 |
fun solve_prem_prem_tac ctxt = |
49429
64ac3471005a
cleaner way of dealing with the set functions of sums and products
blanchet
parents:
49428
diff
changeset
|
143 |
REPEAT o (eresolve_tac @{thms bexE rev_bexI} ORELSE' rtac @{thm rev_bexI[OF UNIV_I]} ORELSE' |
51798 | 144 |
hyp_subst_tac ctxt ORELSE' resolve_tac @{thms disjI1 disjI2}) THEN' |
49429
64ac3471005a
cleaner way of dealing with the set functions of sums and products
blanchet
parents:
49428
diff
changeset
|
145 |
(rtac refl ORELSE' atac ORELSE' rtac @{thm singletonI}); |
49426 | 146 |
|
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
147 |
fun mk_induct_leverage_prem_prems_tac ctxt nn kks fp_abs_inverses abs_inverses set_maps |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
148 |
pre_set_defs = |
52324 | 149 |
HEADGOAL (EVERY' (maps (fn kk => [select_prem_tac nn (dtac meta_spec) kk, etac meta_mp, |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
150 |
SELECT_GOAL (unfold_thms_tac ctxt (pre_set_defs @ fp_abs_inverses @ abs_inverses @ set_maps @ |
55966 | 151 |
sumprod_thms_set)), |
52324 | 152 |
solve_prem_prem_tac ctxt]) (rev kks))); |
49368 | 153 |
|
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
154 |
fun mk_induct_discharge_prem_tac ctxt nn n fp_abs_inverses abs_inverses set_maps pre_set_defs m k |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
155 |
kks = |
49429
64ac3471005a
cleaner way of dealing with the set functions of sums and products
blanchet
parents:
49428
diff
changeset
|
156 |
let val r = length kks in |
52324 | 157 |
HEADGOAL (EVERY' [select_prem_tac n (rotate_tac 1) k, rotate_tac ~1, hyp_subst_tac ctxt, |
158 |
REPEAT_DETERM_N m o (dtac meta_spec THEN' rotate_tac ~1)]) THEN |
|
49391 | 159 |
EVERY [REPEAT_DETERM_N r |
52324 | 160 |
(HEADGOAL (rotate_tac ~1 THEN' dtac meta_mp THEN' rotate_tac 1) THEN prefer_tac 2), |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
54241
diff
changeset
|
161 |
if r > 0 then ALLGOALS (Goal.norm_hhf_tac ctxt) else all_tac, HEADGOAL atac, |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
162 |
mk_induct_leverage_prem_prems_tac ctxt nn kks fp_abs_inverses abs_inverses set_maps |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
163 |
pre_set_defs] |
49391 | 164 |
end; |
49368 | 165 |
|
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
166 |
fun mk_induct_tac ctxt nn ns mss kkss ctr_defs ctor_induct' fp_abs_inverses abs_inverses set_maps |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
167 |
pre_set_defss = |
49590 | 168 |
let val n = Integer.sum ns in |
54922 | 169 |
unfold_thms_tac ctxt ctr_defs THEN HEADGOAL (rtac ctor_induct') THEN |
54923
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
blanchet
parents:
54922
diff
changeset
|
170 |
co_induct_inst_as_projs_tac ctxt 0 THEN |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
171 |
EVERY (map4 (EVERY oooo map3 o |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
172 |
mk_induct_discharge_prem_tac ctxt nn n fp_abs_inverses abs_inverses set_maps) |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
173 |
pre_set_defss mss (unflat mss (1 upto n)) kkss) |
49368 | 174 |
end; |
175 |
||
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
176 |
fun mk_coinduct_same_ctr_tac ctxt rel_eqs pre_rel_def fp_abs_inverse abs_inverse dtor_ctor ctr_def |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
177 |
discs sels = |
51798 | 178 |
hyp_subst_tac ctxt THEN' |
49665 | 179 |
CONVERSION (hhf_concl_conv |
180 |
(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
|
181 |
SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: dtor_ctor :: sels)) THEN' |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
182 |
SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: fp_abs_inverse :: abs_inverse :: dtor_ctor :: |
55966 | 183 |
sels @ sumprod_thms_rel @ @{thms o_apply vimage2p_def})) THEN' |
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49590
diff
changeset
|
184 |
(atac ORELSE' REPEAT o etac conjE THEN' |
56765 | 185 |
full_simp_tac (ss_only (no_refl discs @ rel_eqs @ more_simp_thms) ctxt) THEN' |
54241 | 186 |
REPEAT o etac conjE THEN_MAYBE' REPEAT o hyp_subst_tac ctxt THEN' |
187 |
REPEAT o (resolve_tac [refl, conjI] ORELSE' atac)); |
|
49590 | 188 |
|
52966 | 189 |
fun mk_coinduct_distinct_ctrs_tac ctxt discs discs' = |
54198
4fadf746f2d5
got rid of annoying duplicate rewrite rule warnings
blanchet
parents:
53690
diff
changeset
|
190 |
let |
4fadf746f2d5
got rid of annoying duplicate rewrite rule warnings
blanchet
parents:
53690
diff
changeset
|
191 |
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
|
192 |
|> distinct Thm.eq_thm_prop; |
4fadf746f2d5
got rid of annoying duplicate rewrite rule warnings
blanchet
parents:
53690
diff
changeset
|
193 |
in |
4fadf746f2d5
got rid of annoying duplicate rewrite rule warnings
blanchet
parents:
53690
diff
changeset
|
194 |
hyp_subst_tac ctxt THEN' REPEAT o etac conjE THEN' |
4fadf746f2d5
got rid of annoying duplicate rewrite rule warnings
blanchet
parents:
53690
diff
changeset
|
195 |
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
|
196 |
end; |
49590 | 197 |
|
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
198 |
fun mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn kk n pre_rel_def fp_abs_inverse abs_inverse |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
199 |
dtor_ctor exhaust ctr_defs discss selss = |
49590 | 200 |
let val ks = 1 upto n in |
54837 | 201 |
EVERY' ([rtac allI, rtac allI, rtac impI, select_prem_tac nn (dtac meta_spec) kk, |
54923
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
blanchet
parents:
54922
diff
changeset
|
202 |
dtac meta_spec, dtac meta_mp, atac, rtac exhaust, K (co_induct_inst_as_projs_tac ctxt 0), |
51798 | 203 |
hyp_subst_tac ctxt] @ |
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49590
diff
changeset
|
204 |
map4 (fn k => fn ctr_def => fn discs => fn sels => |
54923
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
blanchet
parents:
54922
diff
changeset
|
205 |
EVERY' ([rtac 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
|
206 |
map2 (fn k' => fn discs' => |
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49590
diff
changeset
|
207 |
if k' = k then |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
208 |
mk_coinduct_same_ctr_tac ctxt rel_eqs' pre_rel_def fp_abs_inverse abs_inverse |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
209 |
dtor_ctor ctr_def discs sels |
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49590
diff
changeset
|
210 |
else |
52966 | 211 |
mk_coinduct_distinct_ctrs_tac ctxt discs discs') ks discss)) ks ctr_defs discss selss) |
49590 | 212 |
end; |
213 |
||
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
214 |
fun mk_coinduct_tac ctxt rel_eqs' nn ns dtor_coinduct' pre_rel_defs fp_abs_inverses abs_inverses |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
215 |
dtor_ctors exhausts ctr_defss discsss selsss = |
52324 | 216 |
HEADGOAL (rtac dtor_coinduct' THEN' |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
217 |
EVERY' (map10 (mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn) |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
218 |
(1 upto nn) ns pre_rel_defs fp_abs_inverses abs_inverses dtor_ctors exhausts ctr_defss discsss |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55642
diff
changeset
|
219 |
selsss)); |
49590 | 220 |
|
57525 | 221 |
fun mk_rel_cases_tac ctxt ct1 ct2 exhaust injects rel_injects distincts rel_distincts = |
222 |
HEADGOAL (rtac (cterm_instantiate_pos [SOME ct1] exhaust) THEN_ALL_NEW |
|
223 |
rtac (cterm_instantiate_pos [SOME ct2] exhaust) THEN_ALL_NEW |
|
224 |
hyp_subst_tac ctxt) THEN |
|
57528 | 225 |
unfold_thms_tac ctxt (injects @ rel_injects @ @{thms conj_imp_eq_imp_imp simp_thms(6) |
57525 | 226 |
True_implies_equals conj_imp_eq_imp_imp} @ |
57529 | 227 |
map (fn thm => thm RS eqFalseI) (distincts @ rel_distincts) @ |
228 |
map (fn thm => thm RS eqTrueI) rel_injects) THEN |
|
57562 | 229 |
TRYALL (atac ORELSE' etac FalseE ORELSE' |
230 |
(REPEAT_DETERM o dtac @{thm meta_spec} THEN' |
|
231 |
TRY o filter_prems_tac |
|
232 |
(forall (curry (op <>) (HOLogic.mk_Trueprop @{term False})) o Logic.strip_imp_prems) THEN' |
|
233 |
REPEAT_DETERM o (dtac @{thm meta_mp} THEN' rtac refl) THEN' Goal.assume_rule_tac ctxt)); |
|
57525 | 234 |
|
57301
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57152
diff
changeset
|
235 |
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
|
236 |
dtor_ctors ctor_injects abs_injects rel_pre_defs abs_inverses nesting_rel_eqs = |
57301
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57152
diff
changeset
|
237 |
rtac dtor_rel_coinduct 1 THEN |
57670 | 238 |
EVERY (map11 (fn ct => fn assm => fn exhaust => fn discs => fn sels => fn ctor_defs => |
239 |
fn dtor_ctor => fn ctor_inject => fn abs_inject => fn rel_pre_def => fn abs_inverse => |
|
240 |
(rtac exhaust THEN_ALL_NEW (rtac exhaust THEN_ALL_NEW |
|
241 |
(dtac (rotate_prems (~1) (cterm_instantiate_pos [NONE, NONE, NONE, NONE, SOME ct] |
|
242 |
@{thm arg_cong2} RS iffD1)) THEN' |
|
243 |
atac THEN' atac THEN' hyp_subst_tac ctxt THEN' dtac assm THEN' |
|
244 |
REPEAT_DETERM o etac conjE))) 1 THEN |
|
57668 | 245 |
unfold_thms_tac ctxt ((discs RL [eqTrueI, eqFalseI]) @ sels @ simp_thms') THEN |
57528 | 246 |
unfold_thms_tac ctxt (dtor_ctor :: rel_pre_def :: abs_inverse :: ctor_inject :: |
57471 | 247 |
abs_inject :: ctor_defs @ nesting_rel_eqs @ simp_thms' @ @{thms BNF_Comp.id_bnf_comp_def |
248 |
rel_sum_simps rel_prod_apply vimage2p_def Inl_Inr_False iffD2[OF eq_False Inr_not_Inl] |
|
249 |
sum.inject prod.inject}) THEN |
|
57301
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57152
diff
changeset
|
250 |
REPEAT_DETERM (HEADGOAL ((REPEAT_DETERM o etac conjE) THEN' (REPEAT_DETERM o rtac conjI) THEN' |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57152
diff
changeset
|
251 |
(rtac refl ORELSE' atac)))) |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57152
diff
changeset
|
252 |
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
|
253 |
abs_inverses); |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57152
diff
changeset
|
254 |
|
57471 | 255 |
fun mk_rel_induct0_tac ctxt ctor_rel_induct assms cterms exhausts ctor_defss ctor_injects |
256 |
rel_pre_list_defs Abs_inverses nesting_rel_eqs = |
|
257 |
rtac ctor_rel_induct 1 THEN EVERY (map6 (fn cterm => fn exhaust => fn ctor_defs => |
|
258 |
fn ctor_inject => fn rel_pre_list_def => fn Abs_inverse => |
|
259 |
HEADGOAL (rtac exhaust THEN_ALL_NEW (rtac exhaust THEN_ALL_NEW |
|
260 |
(rtac (cterm_instantiate_pos (replicate 4 NONE @ [SOME cterm]) @{thm arg_cong2} RS iffD2) |
|
261 |
THEN' atac THEN' atac THEN' TRY o resolve_tac assms))) THEN |
|
57700 | 262 |
unfold_thms_tac ctxt (ctor_inject :: rel_pre_list_def :: ctor_defs @ nesting_rel_eqs @ |
57471 | 263 |
@{thms BNF_Comp.id_bnf_comp_def vimage2p_def}) THEN |
264 |
TRYALL (hyp_subst_tac ctxt) THEN |
|
265 |
unfold_tac ctxt (Abs_inverse :: @{thms rel_sum_simps rel_prod_apply Inl_Inr_False |
|
266 |
Inr_Inl_False sum.inject prod.inject}) THEN |
|
267 |
TRYALL (etac FalseE ORELSE' (REPEAT_DETERM o etac conjE) THEN' atac)) |
|
268 |
cterms exhausts ctor_defss ctor_injects rel_pre_list_defs Abs_inverses); |
|
269 |
||
57563 | 270 |
fun mk_rel_sel_tac ctxt ct1 ct2 exhaust discs sels rel_injects distincts rel_distincts = |
271 |
HEADGOAL (rtac (cterm_instantiate_pos [SOME ct1] exhaust) THEN_ALL_NEW |
|
272 |
rtac (cterm_instantiate_pos [SOME ct2] exhaust) THEN_ALL_NEW |
|
273 |
hyp_subst_tac ctxt) THEN |
|
274 |
Local_Defs.unfold_tac ctxt (sels @ rel_injects @ @{thms simp_thms(6,7,8,11,12,15,16,21,22,24)} @ |
|
275 |
((discs @ distincts) RL @{thms iffD2[OF eq_True] iffD2[OF eq_False]}) @ |
|
276 |
(rel_injects RL @{thms iffD2[OF eq_True]}) @ |
|
277 |
(rel_distincts RL @{thms iffD2[OF eq_False]})) THEN |
|
278 |
TRYALL (resolve_tac [TrueI, refl]); |
|
279 |
||
57046
b3613d7e108b
generate 'sel_map[simp]' theorem for (co)datatypes and tuning 'disc_map_iff'
desharna
parents:
56991
diff
changeset
|
280 |
fun mk_sel_map_tac ctxt ct exhaust discs maps sels = |
b3613d7e108b
generate 'sel_map[simp]' theorem for (co)datatypes and tuning 'disc_map_iff'
desharna
parents:
56991
diff
changeset
|
281 |
TRYALL Goal.conjunction_tac THEN |
b3613d7e108b
generate 'sel_map[simp]' theorem for (co)datatypes and tuning 'disc_map_iff'
desharna
parents:
56991
diff
changeset
|
282 |
ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW |
b3613d7e108b
generate 'sel_map[simp]' theorem for (co)datatypes and tuning 'disc_map_iff'
desharna
parents:
56991
diff
changeset
|
283 |
REPEAT_DETERM o hyp_subst_tac ctxt) THEN |
57529 | 284 |
unfold_thms_tac ctxt ((discs RL [eqTrueI, eqFalseI]) @ |
57046
b3613d7e108b
generate 'sel_map[simp]' theorem for (co)datatypes and tuning 'disc_map_iff'
desharna
parents:
56991
diff
changeset
|
285 |
@{thms not_True_eq_False not_False_eq_True}) THEN |
b3613d7e108b
generate 'sel_map[simp]' theorem for (co)datatypes and tuning 'disc_map_iff'
desharna
parents:
56991
diff
changeset
|
286 |
TRYALL (etac FalseE ORELSE' etac @{thm TrueE}) THEN |
57528 | 287 |
unfold_thms_tac ctxt (maps @ sels) THEN |
57046
b3613d7e108b
generate 'sel_map[simp]' theorem for (co)datatypes and tuning 'disc_map_iff'
desharna
parents:
56991
diff
changeset
|
288 |
ALLGOALS (rtac refl); |
b3613d7e108b
generate 'sel_map[simp]' theorem for (co)datatypes and tuning 'disc_map_iff'
desharna
parents:
56991
diff
changeset
|
289 |
|
57152 | 290 |
fun mk_sel_set_tac ctxt ct exhaust discs sels sets = |
291 |
TRYALL Goal.conjunction_tac THEN |
|
292 |
ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW |
|
293 |
REPEAT_DETERM o hyp_subst_tac ctxt) THEN |
|
57529 | 294 |
unfold_thms_tac ctxt ((discs RL [eqTrueI, eqFalseI]) @ |
57152 | 295 |
@{thms not_True_eq_False not_False_eq_True}) THEN |
296 |
TRYALL (etac FalseE ORELSE' etac @{thm TrueE}) THEN |
|
57528 | 297 |
unfold_thms_tac ctxt (sels @ sets) THEN |
57558
6bb3dd7f8097
took out 'rel_cases' for now because of failing tactic
blanchet
parents:
57529
diff
changeset
|
298 |
ALLGOALS (REPEAT o (resolve_tac @{thms UnI1 UnI2 imageI} ORELSE' |
57152 | 299 |
eresolve_tac @{thms UN_I UN_I[rotated] imageE} ORELSE' |
300 |
hyp_subst_tac ctxt) THEN' |
|
301 |
(rtac @{thm singletonI} ORELSE' atac)); |
|
302 |
||
56956 | 303 |
fun mk_set_empty_tac ctxt ct exhaust sets discs = |
304 |
TRYALL Goal.conjunction_tac THEN |
|
56991
8e9ca31e9b8e
generate 'disc_map_iff[simp]' theorem for (co)datatypes
desharna
parents:
56990
diff
changeset
|
305 |
ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW |
56990
299b026cc5af
fix 'set_empty' theorem when the discriminator is 'op ='
desharna
parents:
56956
diff
changeset
|
306 |
REPEAT_DETERM o hyp_subst_tac ctxt) THEN |
299b026cc5af
fix 'set_empty' theorem when the discriminator is 'op ='
desharna
parents:
56956
diff
changeset
|
307 |
unfold_thms_tac ctxt (sets @ map_filter (fn thm => |
57529 | 308 |
SOME (thm RS eqFalseI) handle THM _ => NONE) discs) THEN |
56991
8e9ca31e9b8e
generate 'disc_map_iff[simp]' theorem for (co)datatypes
desharna
parents:
56990
diff
changeset
|
309 |
ALLGOALS (rtac refl ORELSE' etac FalseE); |
56956 | 310 |
|
57700 | 311 |
fun mk_set_induct0_tac ctxt cts assms dtor_set_inducts exhausts set_pre_defs ctor_defs dtor_ctors |
312 |
Abs_pre_inverses = |
|
313 |
let |
|
314 |
val assms_ctor_defs = |
|
315 |
map (unfold_thms ctxt (@{thm BNF_Comp.id_bnf_comp_def} :: ctor_defs)) assms; |
|
316 |
val exhausts' = map (fn thm => thm RS @{thm asm_rl[of "P x y" for P x y]}) exhausts |
|
317 |
|> map2 (fn ct => cterm_instantiate_pos [NONE, SOME ct]) cts; |
|
318 |
in |
|
319 |
ALLGOALS (resolve_tac dtor_set_inducts) THEN |
|
320 |
TRYALL (resolve_tac exhausts' THEN_ALL_NEW |
|
321 |
(resolve_tac (map (fn ct => refl RS |
|
322 |
cterm_instantiate_pos (replicate 4 NONE @ [SOME ct]) @{thm arg_cong2} RS iffD2) cts) |
|
323 |
THEN' atac THEN' hyp_subst_tac ctxt)) THEN |
|
324 |
unfold_thms_tac ctxt (Abs_pre_inverses @ dtor_ctors @ set_pre_defs @ ctor_defs @ |
|
325 |
@{thms BNF_Comp.id_bnf_comp_def o_apply sum_set_simps prod_set_simps UN_empty UN_insert |
|
326 |
Un_empty_left Un_empty_right empty_iff singleton_iff}) THEN |
|
327 |
REPEAT_DETERM (HEADGOAL |
|
328 |
(TRY o etac UnE THEN' TRY o etac @{thm singletonE} THEN' TRY o hyp_subst_tac ctxt THEN' |
|
329 |
REPEAT_DETERM o eresolve_tac @{thms UN_E UnE singletonE} THEN' |
|
330 |
fold (curry (op ORELSE')) (map (fn thm => |
|
331 |
funpow (length (prems_of thm)) (fn tac => tac THEN' atac) (rtac thm)) assms_ctor_defs) |
|
332 |
(etac FalseE))) |
|
333 |
end; |
|
334 |
||
49123
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
diff
changeset
|
335 |
end; |