2337
+ − 1
(* Title: nominal_dt_alpha.ML
+ − 2
Author: Christian Urban
+ − 3
Author: Cezary Kaliszyk
+ − 4
2595
+ − 5
Performing quotient constructions, lifting theorems and
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 6
deriving support properties for the quotient types.
2337
+ − 7
*)
+ − 8
+ − 9
signature NOMINAL_DT_QUOT =
+ − 10
sig
2400
+ − 11
val define_qtypes: (string list * binding * mixfix) list -> typ list -> term list ->
3045
d0ad264f8c4f
updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 12
thm list -> local_theory -> Quotient_Info.quotients list * local_theory
2337
+ − 13
3157
de89c95c5377
updated to latest changes (10 April) to quotient package (lift_raw_const only takes dummy theorem TrueI....in the future this will not work anymore)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 14
val define_qconsts: typ list -> (string * term * mixfix * thm) list -> local_theory ->
3045
d0ad264f8c4f
updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 15
Quotient_Info.quotconsts list * local_theory
2346
+ − 16
2431
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 17
val define_qperms: typ list -> string list -> (string * sort) list ->
3157
de89c95c5377
updated to latest changes (10 April) to quotient package (lift_raw_const only takes dummy theorem TrueI....in the future this will not work anymore)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 18
(string * term * mixfix * thm) list -> thm list -> local_theory -> local_theory
2400
+ − 19
2431
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 20
val define_qsizes: typ list -> string list -> (string * sort) list ->
3157
de89c95c5377
updated to latest changes (10 April) to quotient package (lift_raw_const only takes dummy theorem TrueI....in the future this will not work anymore)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 21
(string * term * mixfix * thm) list -> local_theory -> local_theory
2426
+ − 22
2434
+ − 23
val lift_thms: typ list -> thm list -> thm list -> Proof.context -> thm list * Proof.context
2475
486d4647bb37
supp-proofs work except for CoreHaskell and Modules (induct is probably not finding the correct instance)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 24
2595
+ − 25
val prove_supports: Proof.context -> thm list -> term list -> thm list
+ − 26
val prove_fsupp: Proof.context -> typ list -> thm -> thm list -> thm list
+ − 27
+ − 28
val fs_instance: typ list -> string list -> (string * sort) list -> thm list ->
+ − 29
local_theory -> local_theory
+ − 30
+ − 31
val prove_fv_supp: typ list -> term list -> term list -> term list -> term list -> thm list ->
+ − 32
thm list -> thm list -> thm list -> thm -> bclause list list -> Proof.context -> thm list
+ − 33
+ − 34
val prove_bns_finite: typ list -> term list -> thm -> thm list -> Proof.context -> thm list
+ − 35
+ − 36
val prove_perm_bn_alpha_thms: typ list -> term list -> term list -> thm -> thm list -> thm list ->
+ − 37
thm list -> Proof.context -> thm list
2598
+ − 38
+ − 39
val prove_permute_bn_thms: typ list -> term list -> term list -> thm -> thm list -> thm list ->
+ − 40
thm list -> Proof.context -> thm list
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 41
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 42
val prove_strong_exhausts: Proof.context -> thm list -> bclause list list list -> thm list ->
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 43
thm list -> thm list -> thm list -> thm list -> thm list
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 44
2630
+ − 45
val prove_strong_induct: Proof.context -> thm -> thm list -> thm list -> bclause list list list ->
+ − 46
thm list
2337
+ − 47
end
+ − 48
+ − 49
structure Nominal_Dt_Quot: NOMINAL_DT_QUOT =
+ − 50
struct
+ − 51
2765
7ac5e5c86c7d
introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 52
open Nominal_Permeq
7ac5e5c86c7d
introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 53
2629
ffb5a181844b
proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 54
fun lookup xs x = the (AList.lookup (op=) xs x)
ffb5a181844b
proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 55
2475
486d4647bb37
supp-proofs work except for CoreHaskell and Modules (induct is probably not finding the correct instance)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 56
2337
+ − 57
(* defines the quotient types *)
2400
+ − 58
fun define_qtypes qtys_descr alpha_tys alpha_trms alpha_equivp_thms lthy =
2476
+ − 59
let
+ − 60
val qty_args1 = map2 (fn ty => fn trm => (ty, trm, false)) alpha_tys alpha_trms
3239
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 61
val qty_args2 = map2 (fn descr => fn args1 => (descr, args1, (NONE, NONE))) qtys_descr qty_args1
3056
756f48abb40a
updated to changes in the quotient package (patch by Ondrej Kuncar)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 62
val qty_args3 = qty_args2 ~~ alpha_equivp_thms
2476
+ − 63
in
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 64
fold_map (Quotient_Type.add_quotient_type {overloaded = false}) qty_args3 lthy
2476
+ − 65
end
2337
+ − 66
3158
+ − 67
(* a wrapper for lifting a raw constant *)
+ − 68
fun lift_raw_const qtys (qconst_name, rconst, mx', rsp_thm) lthy =
+ − 69
let
+ − 70
val rty = fastype_of rconst
+ − 71
val qty = Quotient_Term.derive_qtyp lthy qtys rty
+ − 72
val lhs_raw = Free (qconst_name, qty)
+ − 73
val rhs_raw = rconst
+ − 74
+ − 75
val raw_var = (Binding.name qconst_name, NONE, mx')
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 76
val ((binding, _, mx), ctxt) = Proof_Context.cert_var raw_var lthy
3158
+ − 77
val lhs = Syntax.check_term ctxt lhs_raw
+ − 78
val rhs = Syntax.check_term ctxt rhs_raw
+ − 79
+ − 80
val (lhs_str, lhs_ty) = dest_Free lhs handle TERM _ => error "Constant already defined."
+ − 81
val _ = if null (strip_abs_vars rhs) then () else error "The definiens cannot be an abstraction"
+ − 82
val _ = if is_Const rhs then () else warning "The definiens is not a constant"
+ − 83
+ − 84
in
3245
+ − 85
Quotient_Def.add_quotient_def (((binding, mx), Binding.empty_atts), (lhs, rhs)) rsp_thm ctxt
3158
+ − 86
end
+ − 87
2338
+ − 88
2337
+ − 89
(* defines quotient constants *)
2400
+ − 90
fun define_qconsts qtys consts_specs lthy =
2476
+ − 91
let
+ − 92
val (qconst_infos, lthy') =
3158
+ − 93
fold_map (lift_raw_const qtys) consts_specs lthy
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 94
val phi =
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 95
Proof_Context.export_morphism lthy'
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 96
(Proof_Context.transfer (Proof_Context.theory_of lthy') lthy)
2476
+ − 97
in
3045
d0ad264f8c4f
updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 98
(map (Quotient_Info.transform_quotconsts phi) qconst_infos, lthy')
2476
+ − 99
end
2337
+ − 100
+ − 101
2400
+ − 102
(* defines the quotient permutations and proves pt-class *)
2431
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 103
fun define_qperms qtys qfull_ty_names tvs perm_specs raw_perm_laws lthy =
2476
+ − 104
let
+ − 105
val lthy1 =
+ − 106
lthy
+ − 107
|> Local_Theory.exit_global
+ − 108
|> Class.instantiation (qfull_ty_names, tvs, @{sort pt})
2398
+ − 109
3161
aa1ba91ed1ff
Pass proper rsp theorems for constructors and for size
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
diff
changeset
+ − 110
val (_, lthy2) = define_qconsts qtys perm_specs lthy1
2431
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 111
3046
9b0324e1293b
all examples work again after quotient package has been "de-localised"
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 112
val ((_, raw_perm_laws'), lthy3) = Variable.importT raw_perm_laws lthy2
2398
+ − 113
2476
+ − 114
val lifted_perm_laws =
+ − 115
map (Quotient_Tacs.lifted lthy3 qtys []) raw_perm_laws'
+ − 116
|> Variable.exportT lthy3 lthy2
2398
+ − 117
3239
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 118
fun tac ctxt =
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 119
Class.intro_classes_tac ctxt [] THEN
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 120
(ALLGOALS (resolve_tac ctxt lifted_perm_laws))
2476
+ − 121
in
+ − 122
lthy2
+ − 123
|> Class.prove_instantiation_exit tac
+ − 124
|> Named_Target.theory_init
+ − 125
end
2346
+ − 126
2337
+ − 127
2400
+ − 128
(* defines the size functions and proves size-class *)
2431
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 129
fun define_qsizes qtys qfull_ty_names tvs size_specs lthy =
2476
+ − 130
let
3239
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 131
fun tac ctxt = Class.intro_classes_tac ctxt []
2476
+ − 132
in
+ − 133
lthy
+ − 134
|> Local_Theory.exit_global
+ − 135
|> Class.instantiation (qfull_ty_names, tvs, @{sort size})
+ − 136
|> snd o (define_qconsts qtys size_specs)
+ − 137
|> Class.prove_instantiation_exit tac
+ − 138
|> Named_Target.theory_init
+ − 139
end
2400
+ − 140
2426
+ − 141
2475
486d4647bb37
supp-proofs work except for CoreHaskell and Modules (induct is probably not finding the correct instance)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 142
(* lifts a theorem and cleans all "_raw" parts
2431
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 143
from variable names *)
2426
+ − 144
2431
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 145
local
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 146
val any = Scan.one (Symbol.not_eof)
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 147
val raw = Scan.this_string "_raw"
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 148
val exclude =
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 149
Scan.repeat (Scan.unless raw any) --| raw >> implode
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 150
val parser = Scan.repeat (exclude || any)
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 151
in
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 152
fun unraw_str s =
2574
+ − 153
s |> raw_explode
2476
+ − 154
|> Scan.finite Symbol.stopper parser >> implode
+ − 155
|> fst
2431
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 156
end
2426
+ − 157
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 158
fun unraw_vars_thm ctxt thm =
2476
+ − 159
let
+ − 160
fun unraw_var_str ((s, i), T) = ((unraw_str s, i), T)
2426
+ − 161
3239
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 162
val vars = Term.add_vars (Thm.prop_of thm) []
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 163
val vars' = map (Thm.cterm_of ctxt o Var o unraw_var_str) vars
2476
+ − 164
in
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 165
Thm.instantiate ([], (vars ~~ vars')) thm
2476
+ − 166
end
2426
+ − 167
+ − 168
fun unraw_bounds_thm th =
2476
+ − 169
let
+ − 170
val trm = Thm.prop_of th
+ − 171
val trm' = Term.map_abs_vars unraw_str trm
+ − 172
in
+ − 173
Thm.rename_boundvars trm trm' th
+ − 174
end
2426
+ − 175
2434
+ − 176
fun lift_thms qtys simps thms ctxt =
+ − 177
(map (Quotient_Tacs.lifted ctxt qtys simps
+ − 178
#> unraw_bounds_thm
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 179
#> unraw_vars_thm ctxt
2434
+ − 180
#> Drule.zero_var_indexes) thms, ctxt)
+ − 181
2595
+ − 182
+ − 183
+ − 184
fun mk_supports_goal ctxt qtrm =
+ − 185
let
+ − 186
val vs = fresh_args ctxt qtrm
+ − 187
val rhs = list_comb (qtrm, vs)
+ − 188
val lhs = fold (curry HOLogic.mk_prod) vs @{term "()"}
+ − 189
|> mk_supp
+ − 190
in
+ − 191
mk_supports lhs rhs
+ − 192
|> HOLogic.mk_Trueprop
+ − 193
end
+ − 194
+ − 195
fun supports_tac ctxt perm_simps =
+ − 196
let
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 197
val simpset1 =
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 198
put_simpset HOL_basic_ss ctxt addsimps @{thms supports_def fresh_def[symmetric]}
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 199
val simpset2 =
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 200
put_simpset HOL_ss ctxt addsimps @{thms swap_fresh_fresh fresh_Pair}
2595
+ − 201
in
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 202
EVERY' [ simp_tac simpset1,
2765
7ac5e5c86c7d
introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 203
eqvt_tac ctxt (eqvt_strict_config addpres perm_simps),
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 204
simp_tac simpset2 ]
2595
+ − 205
end
+ − 206
+ − 207
fun prove_supports_single ctxt perm_simps qtrm =
+ − 208
let
+ − 209
val goal = mk_supports_goal ctxt qtrm
+ − 210
val ctxt' = Variable.auto_fixes goal ctxt
+ − 211
in
+ − 212
Goal.prove ctxt' [] [] goal
+ − 213
(K (HEADGOAL (supports_tac ctxt perm_simps)))
3045
d0ad264f8c4f
updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 214
|> singleton (Proof_Context.export ctxt' ctxt)
2595
+ − 215
end
+ − 216
+ − 217
fun prove_supports ctxt perm_simps qtrms =
+ − 218
map (prove_supports_single ctxt perm_simps) qtrms
+ − 219
+ − 220
+ − 221
(* finite supp lemmas for qtypes *)
+ − 222
+ − 223
fun prove_fsupp ctxt qtys qinduct qsupports_thms =
+ − 224
let
+ − 225
val (vs, ctxt') = Variable.variant_fixes (replicate (length qtys) "x") ctxt
+ − 226
val goals = vs ~~ qtys
+ − 227
|> map Free
+ − 228
|> map (mk_finite o mk_supp)
+ − 229
|> foldr1 (HOLogic.mk_conj)
+ − 230
|> HOLogic.mk_Trueprop
+ − 231
+ − 232
val tac =
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 233
EVERY' [ resolve_tac ctxt' @{thms supports_finite},
3239
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 234
resolve_tac ctxt' qsupports_thms,
3227
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 235
asm_simp_tac (put_simpset HOL_ss ctxt'
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 236
addsimps @{thms finite_supp supp_Pair finite_Un}) ]
2595
+ − 237
in
+ − 238
Goal.prove ctxt' [] [] goals
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 239
(K (HEADGOAL (resolve_tac ctxt' [qinduct] THEN_ALL_NEW tac)))
3045
d0ad264f8c4f
updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 240
|> singleton (Proof_Context.export ctxt' ctxt)
3239
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 241
|> Old_Datatype_Aux.split_conj_thm
2595
+ − 242
|> map zero_var_indexes
+ − 243
end
+ − 244
+ − 245
+ − 246
(* finite supp instances *)
+ − 247
+ − 248
fun fs_instance qtys qfull_ty_names tvs qfsupp_thms lthy =
+ − 249
let
+ − 250
val lthy1 =
+ − 251
lthy
+ − 252
|> Local_Theory.exit_global
+ − 253
|> Class.instantiation (qfull_ty_names, tvs, @{sort fs})
+ − 254
3239
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 255
fun tac ctxt =
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 256
Class.intro_classes_tac ctxt [] THEN
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 257
(ALLGOALS (resolve_tac ctxt qfsupp_thms))
2595
+ − 258
in
+ − 259
lthy1
+ − 260
|> Class.prove_instantiation_exit tac
+ − 261
|> Named_Target.theory_init
+ − 262
end
+ − 263
+ − 264
+ − 265
(* proves that fv and fv_bn equals supp *)
+ − 266
+ − 267
fun gen_mk_goals fv supp =
+ − 268
let
+ − 269
val arg_ty =
+ − 270
fastype_of fv
+ − 271
|> domain_type
+ − 272
in
+ − 273
(arg_ty, fn x => HOLogic.mk_eq (fv $ x, supp x))
+ − 274
end
+ − 275
+ − 276
fun mk_fvs_goals fv = gen_mk_goals fv mk_supp
+ − 277
fun mk_fv_bns_goals fv_bn alpha_bn = gen_mk_goals fv_bn (mk_supp_rel alpha_bn)
+ − 278
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 279
fun add_simpset ctxt thms = put_simpset HOL_basic_ss ctxt addsimps thms
2595
+ − 280
+ − 281
fun symmetric thms =
+ − 282
map (fn thm => thm RS @{thm sym}) thms
+ − 283
+ − 284
val supp_Abs_set = @{thms supp_Abs(1)[symmetric]}
+ − 285
val supp_Abs_res = @{thms supp_Abs(2)[symmetric]}
+ − 286
val supp_Abs_lst = @{thms supp_Abs(3)[symmetric]}
+ − 287
+ − 288
fun mk_supp_abs ctxt (BC (Set, _, _)) = EqSubst.eqsubst_tac ctxt [1] supp_Abs_set
+ − 289
| mk_supp_abs ctxt (BC (Res, _, _)) = EqSubst.eqsubst_tac ctxt [1] supp_Abs_res
+ − 290
| mk_supp_abs ctxt (BC (Lst, _, _)) = EqSubst.eqsubst_tac ctxt [1] supp_Abs_lst
+ − 291
+ − 292
fun mk_supp_abs_tac ctxt [] = []
+ − 293
| mk_supp_abs_tac ctxt (BC (_, [], _)::xs) = mk_supp_abs_tac ctxt xs
+ − 294
| mk_supp_abs_tac ctxt (bc::xs) = (DETERM o mk_supp_abs ctxt bc)::mk_supp_abs_tac ctxt xs
+ − 295
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 296
fun mk_bn_supp_abs_tac ctxt trm =
2595
+ − 297
trm
+ − 298
|> fastype_of
+ − 299
|> body_type
+ − 300
|> (fn ty => case ty of
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 301
@{typ "atom set"} => simp_tac (add_simpset ctxt supp_Abs_set)
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 302
| @{typ "atom list"} => simp_tac (add_simpset ctxt supp_Abs_lst)
2595
+ − 303
| _ => raise TERM ("mk_bn_supp_abs_tac", [trm]))
+ − 304
+ − 305
+ − 306
val thms1 = @{thms supp_Pair supp_eqvt[symmetric] Un_assoc conj_assoc}
+ − 307
val thms2 = @{thms de_Morgan_conj Collect_disj_eq finite_Un}
3239
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 308
val thms3 = @{thms alphas prod_alpha_def prod_fv.simps rel_prod_conv permute_prod_def
3229
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 309
prod.rec prod.case prod.inject not_True_eq_False empty_def[symmetric] finite.emptyI}
2595
+ − 310
+ − 311
fun prove_fv_supp qtys qtrms fvs fv_bns alpha_bns fv_simps eq_iffs perm_simps
+ − 312
fv_bn_eqvts qinduct bclausess ctxt =
+ − 313
let
+ − 314
val goals1 = map mk_fvs_goals fvs
+ − 315
val goals2 = map2 mk_fv_bns_goals fv_bns alpha_bns
+ − 316
+ − 317
fun tac ctxt =
+ − 318
SUBGOAL (fn (goal, i) =>
+ − 319
let
+ − 320
val (fv_fun, arg) =
+ − 321
goal |> Envir.eta_contract
+ − 322
|> Logic.strip_assums_concl
+ − 323
|> HOLogic.dest_Trueprop
+ − 324
|> fst o HOLogic.dest_eq
+ − 325
|> dest_comb
+ − 326
val supp_abs_tac =
+ − 327
case (AList.lookup (op=) (qtrms ~~ bclausess) (head_of arg)) of
+ − 328
SOME bclauses => EVERY' (mk_supp_abs_tac ctxt bclauses)
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 329
| NONE => mk_bn_supp_abs_tac ctxt fv_fun
2765
7ac5e5c86c7d
introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 330
val eqvt_rconfig = eqvt_relaxed_config addpres (perm_simps @ fv_bn_eqvts)
2595
+ − 331
in
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 332
EVERY' [ TRY o asm_full_simp_tac (add_simpset ctxt (@{thm supp_Pair[symmetric]}::fv_simps)),
2595
+ − 333
TRY o supp_abs_tac,
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 334
TRY o simp_tac (add_simpset ctxt @{thms supp_def supp_rel_def}),
2765
7ac5e5c86c7d
introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 335
TRY o eqvt_tac ctxt eqvt_rconfig,
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 336
TRY o simp_tac (add_simpset ctxt (@{thms Abs_eq_iff} @ eq_iffs)),
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 337
TRY o asm_full_simp_tac (add_simpset ctxt thms3),
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 338
TRY o simp_tac (add_simpset ctxt thms2),
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 339
TRY o asm_full_simp_tac (add_simpset ctxt (thms1 @ (symmetric fv_bn_eqvts)))] i
2595
+ − 340
end)
+ − 341
in
+ − 342
induct_prove qtys (goals1 @ goals2) qinduct tac ctxt
3226
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 343
|> map (atomize ctxt)
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 344
|> map (simplify (put_simpset HOL_basic_ss ctxt addsimps @{thms fun_eq_iff[symmetric]}))
2595
+ − 345
end
+ − 346
+ − 347
+ − 348
fun prove_bns_finite qtys qbns qinduct qbn_simps ctxt =
+ − 349
let
+ − 350
fun mk_goal qbn =
+ − 351
let
+ − 352
val arg_ty = domain_type (fastype_of qbn)
+ − 353
val finite = @{term "finite :: atom set => bool"}
+ − 354
in
+ − 355
(arg_ty, fn x => finite $ (to_set (qbn $ x)))
+ − 356
end
+ − 357
+ − 358
val props = map mk_goal qbns
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 359
val ss_tac = asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps (qbn_simps @
3229
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 360
@{thms set_simps set_append finite_insert finite.emptyI finite_Un}))
2595
+ − 361
in
+ − 362
induct_prove qtys props qinduct (K ss_tac) ctxt
+ − 363
end
+ − 364
2598
+ − 365
2595
+ − 366
fun prove_perm_bn_alpha_thms qtys qperm_bns alpha_bns qinduct qperm_bn_simps qeq_iffs qalpha_refls ctxt =
+ − 367
let
+ − 368
val ([p], ctxt') = Variable.variant_fixes ["p"] ctxt
+ − 369
val p = Free (p, @{typ perm})
+ − 370
+ − 371
fun mk_goal qperm_bn alpha_bn =
+ − 372
let
+ − 373
val arg_ty = domain_type (fastype_of alpha_bn)
+ − 374
in
+ − 375
(arg_ty, fn x => (mk_id (Abs ("", arg_ty, alpha_bn $ Bound 0 $ (qperm_bn $ p $ Bound 0)))) $ x)
+ − 376
end
+ − 377
+ − 378
val props = map2 mk_goal qperm_bns alpha_bns
+ − 379
val ss = @{thm id_def}::qperm_bn_simps @ qeq_iffs @ qalpha_refls
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 380
val ss_tac = asm_full_simp_tac (put_simpset HOL_ss ctxt addsimps ss)
2595
+ − 381
in
+ − 382
induct_prove qtys props qinduct (K ss_tac) ctxt'
3045
d0ad264f8c4f
updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 383
|> Proof_Context.export ctxt' ctxt
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 384
|> map (simplify (put_simpset HOL_basic_ss ctxt addsimps @{thms id_def}))
2595
+ − 385
end
+ − 386
2598
+ − 387
fun prove_permute_bn_thms qtys qbns qperm_bns qinduct qperm_bn_simps qbn_defs qbn_eqvts ctxt =
+ − 388
let
+ − 389
val ([p], ctxt') = Variable.variant_fixes ["p"] ctxt
+ − 390
val p = Free (p, @{typ perm})
2595
+ − 391
2598
+ − 392
fun mk_goal qbn qperm_bn =
+ − 393
let
+ − 394
val arg_ty = domain_type (fastype_of qbn)
+ − 395
in
+ − 396
(arg_ty, fn x =>
+ − 397
(mk_id (Abs ("", arg_ty,
+ − 398
HOLogic.mk_eq (mk_perm p (qbn $ Bound 0), qbn $ (qperm_bn $ p $ Bound 0)))) $ x))
+ − 399
end
+ − 400
+ − 401
val props = map2 mk_goal qbns qperm_bns
+ − 402
val ss = @{thm id_def}::qperm_bn_simps @ qbn_defs
+ − 403
val ss_tac =
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 404
EVERY' [asm_full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps ss),
2765
7ac5e5c86c7d
introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 405
TRY o eqvt_tac ctxt' (eqvt_strict_config addpres qbn_eqvts),
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 406
TRY o asm_full_simp_tac (put_simpset HOL_basic_ss ctxt')]
2598
+ − 407
in
+ − 408
induct_prove qtys props qinduct (K ss_tac) ctxt'
3045
d0ad264f8c4f
updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 409
|> Proof_Context.export ctxt' ctxt
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 410
|> map (simplify (put_simpset HOL_basic_ss ctxt addsimps @{thms id_def}))
2598
+ − 411
end
2595
+ − 412
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 413
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 414
2629
ffb5a181844b
proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 415
(*** proves strong exhauts theorems ***)
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 416
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 417
(* fixme: move into nominal_library *)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 418
fun abs_const bmode ty =
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 419
let
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 420
val (const_name, binder_ty, abs_ty) =
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 421
case bmode of
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 422
Lst => (@{const_name "Abs_lst"}, @{typ "atom list"}, @{type_name abs_lst})
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 423
| Set => (@{const_name "Abs_set"}, @{typ "atom set"}, @{type_name abs_set})
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 424
| Res => (@{const_name "Abs_res"}, @{typ "atom set"}, @{type_name abs_res})
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 425
in
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 426
Const (const_name, [binder_ty, ty] ---> Type (abs_ty, [ty]))
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 427
end
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 428
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 429
fun mk_abs bmode trm1 trm2 =
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 430
abs_const bmode (fastype_of trm2) $ trm1 $ trm2
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 431
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 432
fun is_abs_eq thm =
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 433
let
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 434
fun is_abs trm =
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 435
case (head_of trm) of
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 436
Const (@{const_name "Abs_set"}, _) => true
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 437
| Const (@{const_name "Abs_lst"}, _) => true
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 438
| Const (@{const_name "Abs_res"}, _) => true
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 439
| _ => false
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 440
in
3239
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 441
thm |> Thm.prop_of
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 442
|> HOLogic.dest_Trueprop
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 443
|> HOLogic.dest_eq
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 444
|> fst
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 445
|> is_abs
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 446
end
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 447
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 448
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 449
(* adds a freshness condition to the assumptions *)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 450
fun mk_ecase_prems lthy c (params, prems, concl) bclauses =
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 451
let
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 452
val tys = map snd params
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 453
val binders = get_all_binders bclauses
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 454
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 455
fun prep_binder (opt, i) =
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 456
let
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 457
val t = Bound (length tys - i - 1)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 458
in
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 459
case opt of
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 460
NONE => setify_ty lthy (nth tys i) t
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 461
| SOME bn => to_set_ty (fastype_of1 (tys, bn $ t)) (bn $ t)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 462
end
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 463
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 464
val prems' =
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 465
case binders of
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 466
[] => prems (* case: no binders *)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 467
| _ => binders (* case: binders *)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 468
|> map prep_binder
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 469
|> fold_union_env tys
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 470
|> (fn t => mk_fresh_star t c)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 471
|> (fn t => HOLogic.mk_Trueprop t :: prems)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 472
in
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 473
mk_full_horn params prems' concl
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 474
end
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 475
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 476
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 477
(* derives the freshness theorem that there exists a p, such that
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 478
(p o as) #* (c, t1,..., tn) *)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 479
fun fresh_thm ctxt c parms binders bn_finite_thms =
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 480
let
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 481
fun prep_binder (opt, i) =
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 482
case opt of
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 483
NONE => setify ctxt (nth parms i)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 484
| SOME bn => to_set (bn $ (nth parms i))
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 485
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 486
fun prep_binder2 (opt, i) =
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 487
case opt of
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 488
NONE => atomify ctxt (nth parms i)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 489
| SOME bn => bn $ (nth parms i)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 490
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 491
val rhs = HOLogic.mk_tuple ([c] @ parms @ (map prep_binder2 binders))
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 492
val lhs = binders
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 493
|> map prep_binder
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 494
|> fold_union
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 495
|> mk_perm (Bound 0)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 496
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 497
val goal = mk_fresh_star lhs rhs
2637
3890483c674f
final version of the ESOP paper; used set+ instead of res as requested by one reviewer
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 498
|> mk_exists ("p", @{typ perm})
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 499
|> HOLogic.mk_Trueprop
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 500
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 501
val ss = bn_finite_thms @ @{thms supp_Pair finite_supp finite_sets_supp}
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 502
@ @{thms finite.intros finite_Un finite_set finite_fset}
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 503
in
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 504
Goal.prove ctxt [] [] goal
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 505
(K (HEADGOAL (resolve_tac ctxt @{thms at_set_avoiding1}
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 506
THEN_ALL_NEW (simp_tac (put_simpset HOL_ss ctxt addsimps ss)))))
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 507
end
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 508
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 509
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 510
(* derives an abs_eq theorem of the form
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 511
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 512
Exists q. [as].x = [p o as].(q o x) for non-recursive binders
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 513
Exists q. [as].x = [q o as].(q o x) for recursive binders
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 514
*)
3060
+ − 515
fun abs_eq_thm ctxt fprops p parms bn_eqvt permute_bns (bclause as (BC (bmode, binders, bodies))) =
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 516
case binders of
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 517
[] => []
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 518
| _ =>
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 519
let
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 520
val rec_flag = is_recursive_binder bclause
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 521
val binder_trm = comb_binders ctxt bmode parms binders
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 522
val body_trm = foldl1 HOLogic.mk_prod (map (nth parms) bodies)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 523
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 524
val abs_lhs = mk_abs bmode binder_trm body_trm
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 525
val abs_rhs =
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 526
if rec_flag
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 527
then mk_abs bmode (mk_perm (Bound 0) binder_trm) (mk_perm (Bound 0) body_trm)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 528
else mk_abs bmode (mk_perm p binder_trm) (mk_perm (Bound 0) body_trm)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 529
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 530
val abs_eq = HOLogic.mk_eq (abs_lhs, abs_rhs)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 531
val peq = HOLogic.mk_eq (mk_perm (Bound 0) binder_trm, mk_perm p binder_trm)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 532
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 533
val goal = HOLogic.mk_conj (abs_eq, peq)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 534
|> (fn t => HOLogic.mk_exists ("q", @{typ "perm"}, t))
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 535
|> HOLogic.mk_Trueprop
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 536
3229
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 537
val ss = fprops @ @{thms set_simps set_append union_eqvt}
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 538
@ @{thms fresh_star_Un fresh_star_Pair fresh_star_list fresh_star_singleton fresh_star_fset
3060
+ − 539
fresh_star_set}
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 540
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 541
val tac1 =
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 542
if rec_flag
3239
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 543
then resolve_tac ctxt @{thms Abs_rename_set' Abs_rename_res' Abs_rename_lst'}
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 544
else resolve_tac ctxt @{thms Abs_rename_set Abs_rename_res Abs_rename_lst}
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 545
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 546
val tac2 =
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 547
EVERY' [simp_tac (put_simpset HOL_basic_ss ctxt addsimps ss),
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 548
TRY o simp_tac (put_simpset HOL_ss ctxt)]
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 549
in
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 550
[ Goal.prove ctxt [] [] goal (K (HEADGOAL (tac1 THEN_ALL_NEW tac2)))
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 551
|> (if rec_flag
2765
7ac5e5c86c7d
introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 552
then eqvt_rule ctxt (eqvt_strict_config addpres bn_eqvt)
7ac5e5c86c7d
introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 553
else eqvt_rule ctxt (eqvt_strict_config addpres permute_bns)) ]
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 554
end
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 555
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 556
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 557
val setify = @{lemma "xs = ys ==> set xs = set ys" by simp}
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 558
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 559
fun case_tac ctxt c bn_finite_thms eq_iff_thms bn_eqvt permute_bns perm_bn_alphas
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 560
prems bclausess qexhaust_thm =
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 561
let
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 562
fun aux_tac prem bclauses =
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 563
case (get_all_binders bclauses) of
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 564
[] => EVERY' [resolve_tac ctxt [prem], assume_tac ctxt]
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 565
| binders => Subgoal.SUBPROOF (fn {params, prems, concl, context = ctxt, ...} =>
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 566
let
3239
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 567
val parms = map (Thm.term_of o snd) params
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 568
val fthm = fresh_thm ctxt c parms binders bn_finite_thms
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 569
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 570
val ss = @{thms fresh_star_Pair union_eqvt fresh_star_Un}
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 571
val (([(_, fperm)], fprops), ctxt') = Obtain.result
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 572
(fn ctxt' => EVERY1 [eresolve_tac ctxt [exE],
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 573
full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps ss),
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 574
REPEAT o (eresolve_tac ctxt @{thms conjE})]) [fthm] ctxt
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 575
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 576
val abs_eq_thms = flat
3239
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 577
(map (abs_eq_thm ctxt' fprops (Thm.term_of fperm) parms bn_eqvt permute_bns) bclauses)
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 578
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 579
val ((_, eqs), ctxt'') = Obtain.result
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 580
(fn ctxt'' => EVERY1
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 581
[ REPEAT o (eresolve_tac ctxt @{thms exE}),
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 582
REPEAT o (eresolve_tac ctxt @{thms conjE}),
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 583
REPEAT o (dresolve_tac ctxt [setify]),
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 584
full_simp_tac (put_simpset HOL_basic_ss ctxt''
3229
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 585
addsimps @{thms set_append set_simps})]) abs_eq_thms ctxt'
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 586
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 587
val (abs_eqs, peqs) = split_filter is_abs_eq eqs
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 588
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 589
val fprops' =
3227
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 590
map (eqvt_rule ctxt'' (eqvt_strict_config addpres permute_bns)) fprops
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 591
@ map (eqvt_rule ctxt'' (eqvt_strict_config addpres bn_eqvt)) fprops
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 592
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 593
(* for freshness conditions *)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 594
val tac1 = SOLVED' (EVERY'
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 595
[ simp_tac (put_simpset HOL_basic_ss ctxt'' addsimps peqs),
3226
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 596
rewrite_goal_tac ctxt'' (@{thms fresh_star_Un[THEN eq_reflection]}),
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 597
conj_tac ctxt'' (DETERM o resolve_tac ctxt'' fprops') ])
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 598
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 599
(* for equalities between constructors *)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 600
val tac2 = SOLVED' (EVERY'
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 601
[ resolve_tac ctxt [@{thm ssubst} OF prems],
3226
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 602
rewrite_goal_tac ctxt'' (map safe_mk_equiv eq_iff_thms),
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 603
rewrite_goal_tac ctxt'' (map safe_mk_equiv abs_eqs),
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 604
conj_tac ctxt'' (DETERM o resolve_tac ctxt'' (@{thms refl} @ perm_bn_alphas)) ])
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 605
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 606
(* proves goal "P" *)
3239
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 607
val side_thm = Goal.prove ctxt'' [] [] (Thm.term_of concl)
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 608
(K (EVERY1 [ resolve_tac ctxt'' [prem], RANGE [tac1, tac2] ]))
3045
d0ad264f8c4f
updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 609
|> singleton (Proof_Context.export ctxt'' ctxt)
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 610
in
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 611
resolve_tac ctxt [side_thm] 1
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 612
end) ctxt
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 613
in
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 614
EVERY1 [resolve_tac ctxt [qexhaust_thm], RANGE (map2 aux_tac prems bclausess)]
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 615
end
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 616
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 617
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 618
fun prove_strong_exhausts lthy exhausts bclausesss bn_finite_thms eq_iff_thms bn_eqvt permute_bns
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 619
perm_bn_alphas =
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 620
let
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 621
val ((_, exhausts'), lthy') = Variable.import true exhausts lthy
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 622
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 623
val ([c, a], lthy'') = Variable.variant_fixes ["c", "'a"] lthy'
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 624
val c = Free (c, TFree (a, @{sort fs}))
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 625
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 626
val (ecases, main_concls) = exhausts' (* ecases are of the form (params, prems, concl) *)
3239
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 627
|> map Thm.prop_of
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 628
|> map Logic.strip_horn
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 629
|> split_list
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 630
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 631
val ecases' = (map o map) strip_full_horn ecases
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 632
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 633
val premss = (map2 o map2) (mk_ecase_prems lthy'' c) ecases' bclausesss
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 634
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 635
fun tac bclausess exhaust {prems, context} =
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 636
case_tac context c bn_finite_thms eq_iff_thms bn_eqvt permute_bns perm_bn_alphas
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 637
prems bclausess exhaust
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 638
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 639
fun prove prems bclausess exhaust concl =
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 640
Goal.prove lthy'' [] prems concl (tac bclausess exhaust)
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 641
in
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 642
map4 prove premss bclausesss exhausts' main_concls
3045
d0ad264f8c4f
updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 643
|> Proof_Context.export lthy'' lthy
2626
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 644
end
d1bdc281be2b
moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 645
2628
+ − 646
+ − 647
+ − 648
(** strong induction theorems **)
+ − 649
+ − 650
fun add_c_prop c c_ty trm =
+ − 651
let
+ − 652
val (P, arg) = dest_comb trm
+ − 653
val (P_name, P_ty) = dest_Free P
+ − 654
val (ty_args, bool) = strip_type P_ty
+ − 655
in
+ − 656
Free (P_name, (c_ty :: ty_args) ---> bool) $ c $ arg
+ − 657
end
+ − 658
+ − 659
fun add_qnt_c_prop c_name c_ty trm =
+ − 660
trm |> HOLogic.dest_Trueprop
+ − 661
|> incr_boundvars 1
+ − 662
|> add_c_prop (Bound 0) c_ty
+ − 663
|> HOLogic.mk_Trueprop
+ − 664
|> mk_all (c_name, c_ty)
+ − 665
2629
ffb5a181844b
proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 666
fun prep_prem lthy c c_name c_ty bclauses (params, prems, concl) =
2628
+ − 667
let
+ − 668
val tys = map snd params
+ − 669
val binders = get_all_binders bclauses
+ − 670
+ − 671
fun prep_binder (opt, i) =
+ − 672
let
+ − 673
val t = Bound (length tys - i - 1)
+ − 674
in
+ − 675
case opt of
+ − 676
NONE => setify_ty lthy (nth tys i) t
+ − 677
| SOME bn => to_set_ty (fastype_of1 (tys, bn $ t)) (bn $ t)
+ − 678
end
+ − 679
+ − 680
val prems' = prems
+ − 681
|> map (incr_boundvars 1)
+ − 682
|> map (add_qnt_c_prop c_name c_ty)
+ − 683
+ − 684
val prems'' =
+ − 685
case binders of
+ − 686
[] => prems' (* case: no binders *)
+ − 687
| _ => binders (* case: binders *)
+ − 688
|> map prep_binder
+ − 689
|> fold_union_env tys
+ − 690
|> incr_boundvars 1
+ − 691
|> (fn t => mk_fresh_star_ty c_ty t (Bound 0))
+ − 692
|> (fn t => HOLogic.mk_Trueprop t :: prems')
+ − 693
+ − 694
val concl' = concl
+ − 695
|> HOLogic.dest_Trueprop
+ − 696
|> incr_boundvars 1
+ − 697
|> add_c_prop (Bound 0) c_ty
+ − 698
|> HOLogic.mk_Trueprop
+ − 699
in
+ − 700
mk_full_horn (params @ [(c_name, c_ty)]) prems'' concl'
+ − 701
end
+ − 702
2630
+ − 703
fun prove_strong_induct lthy induct exhausts size_thms bclausesss =
2628
+ − 704
let
2635
+ − 705
val ((_, [induct']), lthy') = Variable.import true [induct] lthy
2628
+ − 706
+ − 707
val ([c_name, a], lthy'') = Variable.variant_fixes ["c", "'a"] lthy'
+ − 708
val c_ty = TFree (a, @{sort fs})
+ − 709
val c = Free (c_name, c_ty)
+ − 710
+ − 711
val (prems, concl) = induct'
3239
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 712
|> Thm.prop_of
2628
+ − 713
|> Logic.strip_horn
+ − 714
+ − 715
val concls = concl
+ − 716
|> HOLogic.dest_Trueprop
+ − 717
|> HOLogic.dest_conj
+ − 718
|> map (add_c_prop c c_ty)
+ − 719
|> map HOLogic.mk_Trueprop
+ − 720
+ − 721
val prems' = prems
+ − 722
|> map strip_full_horn
2629
ffb5a181844b
proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 723
|> map2 (prep_prem lthy'' c c_name c_ty) (flat bclausesss)
2628
+ − 724
2629
ffb5a181844b
proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 725
fun pat_tac ctxt thm =
3239
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 726
Subgoal.FOCUS (fn {params, context = ctxt', ...} =>
2629
ffb5a181844b
proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 727
let
3239
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 728
val ty_parms = map (fn (_, ct) => (fastype_of (Thm.term_of ct), ct)) params
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 729
val vs = Term.add_vars (Thm.prop_of thm) []
2629
ffb5a181844b
proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 730
val vs_tys = map (Type.legacy_freeze_type o snd) vs
ffb5a181844b
proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 731
val assigns = map (lookup ty_parms) vs_tys
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 732
val thm' = infer_instantiate ctxt' (map #1 vs ~~ assigns) thm
2629
ffb5a181844b
proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 733
in
3244
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 734
resolve_tac ctxt' [thm'] 1
2629
ffb5a181844b
proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 735
end) ctxt
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 736
THEN_ALL_NEW asm_full_simp_tac (put_simpset HOL_basic_ss ctxt)
2629
ffb5a181844b
proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 737
3227
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 738
fun size_simp_tac ctxt =
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 739
simp_tac (put_simpset size_ss ctxt addsimps (@{thms comp_def snd_conv} @ size_thms))
2630
+ − 740
in
3239
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 741
Goal.prove_common lthy'' NONE [] prems' concls
3227
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 742
(fn {prems, context = ctxt} =>
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 743
Induction_Schema.induction_schema_tac ctxt prems
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 744
THEN RANGE (map (pat_tac ctxt) exhausts) 1
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 745
THEN prove_termination_ind ctxt 1
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 746
THEN ALLGOALS (size_simp_tac ctxt))
3045
d0ad264f8c4f
updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 747
|> Proof_Context.export lthy'' lthy
2628
+ − 748
end
+ − 749
+ − 750
2337
+ − 751
end (* structure *)
+ − 752