597
+ − 1
theory QuotMain
648
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2
imports QuotScript Prove
597
+ − 3
uses ("quotient_info.ML")
+ − 4
("quotient.ML")
+ − 5
("quotient_def.ML")
+ − 6
begin
+ − 7
+ − 8
+ − 9
locale QUOT_TYPE =
+ − 10
fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+ − 11
and Abs :: "('a \<Rightarrow> bool) \<Rightarrow> 'b"
+ − 12
and Rep :: "'b \<Rightarrow> ('a \<Rightarrow> bool)"
+ − 13
assumes equivp: "equivp R"
+ − 14
and rep_prop: "\<And>y. \<exists>x. Rep y = R x"
+ − 15
and rep_inverse: "\<And>x. Abs (Rep x) = x"
+ − 16
and abs_inverse: "\<And>x. (Rep (Abs (R x))) = (R x)"
+ − 17
and rep_inject: "\<And>x y. (Rep x = Rep y) = (x = y)"
+ − 18
begin
+ − 19
+ − 20
definition
+ − 21
ABS::"'a \<Rightarrow> 'b"
+ − 22
where
+ − 23
"ABS x \<equiv> Abs (R x)"
+ − 24
+ − 25
definition
+ − 26
REP::"'b \<Rightarrow> 'a"
+ − 27
where
+ − 28
"REP a = Eps (Rep a)"
+ − 29
+ − 30
lemma lem9:
+ − 31
shows "R (Eps (R x)) = R x"
+ − 32
proof -
+ − 33
have a: "R x x" using equivp by (simp add: equivp_reflp_symp_transp reflp_def)
+ − 34
then have "R x (Eps (R x))" by (rule someI)
+ − 35
then show "R (Eps (R x)) = R x"
+ − 36
using equivp unfolding equivp_def by simp
+ − 37
qed
+ − 38
+ − 39
theorem thm10:
+ − 40
shows "ABS (REP a) \<equiv> a"
+ − 41
apply (rule eq_reflection)
+ − 42
unfolding ABS_def REP_def
+ − 43
proof -
+ − 44
from rep_prop
+ − 45
obtain x where eq: "Rep a = R x" by auto
+ − 46
have "Abs (R (Eps (Rep a))) = Abs (R (Eps (R x)))" using eq by simp
+ − 47
also have "\<dots> = Abs (R x)" using lem9 by simp
+ − 48
also have "\<dots> = Abs (Rep a)" using eq by simp
+ − 49
also have "\<dots> = a" using rep_inverse by simp
+ − 50
finally
+ − 51
show "Abs (R (Eps (Rep a))) = a" by simp
+ − 52
qed
+ − 53
+ − 54
lemma REP_refl:
+ − 55
shows "R (REP a) (REP a)"
+ − 56
unfolding REP_def
+ − 57
by (simp add: equivp[simplified equivp_def])
+ − 58
+ − 59
lemma lem7:
+ − 60
shows "(R x = R y) = (Abs (R x) = Abs (R y))"
+ − 61
apply(rule iffI)
+ − 62
apply(simp)
+ − 63
apply(drule rep_inject[THEN iffD2])
+ − 64
apply(simp add: abs_inverse)
+ − 65
done
+ − 66
+ − 67
theorem thm11:
+ − 68
shows "R r r' = (ABS r = ABS r')"
+ − 69
unfolding ABS_def
+ − 70
by (simp only: equivp[simplified equivp_def] lem7)
+ − 71
+ − 72
+ − 73
lemma REP_ABS_rsp:
+ − 74
shows "R f (REP (ABS g)) = R f g"
+ − 75
and "R (REP (ABS g)) f = R g f"
+ − 76
by (simp_all add: thm10 thm11)
+ − 77
+ − 78
lemma Quotient:
+ − 79
"Quotient R ABS REP"
+ − 80
apply(unfold Quotient_def)
+ − 81
apply(simp add: thm10)
+ − 82
apply(simp add: REP_refl)
+ − 83
apply(subst thm11[symmetric])
+ − 84
apply(simp add: equivp[simplified equivp_def])
+ − 85
done
+ − 86
+ − 87
lemma R_trans:
+ − 88
assumes ab: "R a b"
+ − 89
and bc: "R b c"
+ − 90
shows "R a c"
+ − 91
proof -
+ − 92
have tr: "transp R" using equivp equivp_reflp_symp_transp[of R] by simp
+ − 93
moreover have ab: "R a b" by fact
+ − 94
moreover have bc: "R b c" by fact
+ − 95
ultimately show "R a c" unfolding transp_def by blast
+ − 96
qed
+ − 97
+ − 98
lemma R_sym:
+ − 99
assumes ab: "R a b"
+ − 100
shows "R b a"
+ − 101
proof -
+ − 102
have re: "symp R" using equivp equivp_reflp_symp_transp[of R] by simp
+ − 103
then show "R b a" using ab unfolding symp_def by blast
+ − 104
qed
+ − 105
+ − 106
lemma R_trans2:
+ − 107
assumes ac: "R a c"
+ − 108
and bd: "R b d"
+ − 109
shows "R a b = R c d"
+ − 110
using ac bd
+ − 111
by (blast intro: R_trans R_sym)
+ − 112
+ − 113
lemma REPS_same:
+ − 114
shows "R (REP a) (REP b) \<equiv> (a = b)"
+ − 115
proof -
+ − 116
have "R (REP a) (REP b) = (a = b)"
+ − 117
proof
+ − 118
assume as: "R (REP a) (REP b)"
+ − 119
from rep_prop
+ − 120
obtain x y
+ − 121
where eqs: "Rep a = R x" "Rep b = R y" by blast
+ − 122
from eqs have "R (Eps (R x)) (Eps (R y))" using as unfolding REP_def by simp
+ − 123
then have "R x (Eps (R y))" using lem9 by simp
+ − 124
then have "R (Eps (R y)) x" using R_sym by blast
+ − 125
then have "R y x" using lem9 by simp
+ − 126
then have "R x y" using R_sym by blast
+ − 127
then have "ABS x = ABS y" using thm11 by simp
+ − 128
then have "Abs (Rep a) = Abs (Rep b)" using eqs unfolding ABS_def by simp
+ − 129
then show "a = b" using rep_inverse by simp
+ − 130
next
+ − 131
assume ab: "a = b"
+ − 132
have "reflp R" using equivp equivp_reflp_symp_transp[of R] by simp
+ − 133
then show "R (REP a) (REP b)" unfolding reflp_def using ab by auto
+ − 134
qed
+ − 135
then show "R (REP a) (REP b) \<equiv> (a = b)" by simp
+ − 136
qed
+ − 137
+ − 138
end
+ − 139
+ − 140
section {* type definition for the quotient type *}
+ − 141
+ − 142
(* the auxiliary data for the quotient types *)
+ − 143
use "quotient_info.ML"
+ − 144
+ − 145
declare [[map "fun" = (fun_map, fun_rel)]]
615
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 146
613
+ − 147
(* Throws now an exception: *)
612
+ − 148
(* declare [[map "option" = (bla, blu)]] *)
+ − 149
648
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 150
lemmas [quot_thm] = fun_quotient
597
+ − 151
648
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 152
lemmas [quot_respect] = quot_rel_rsp
597
+ − 153
+ − 154
(* fun_map is not here since equivp is not true *)
+ − 155
(* TODO: option, ... *)
648
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 156
lemmas [quot_equiv] = identity_equivp
597
+ − 157
+ − 158
(* definition of the quotient types *)
+ − 159
(* FIXME: should be called quotient_typ.ML *)
+ − 160
use "quotient.ML"
+ − 161
+ − 162
(* lifting of constants *)
+ − 163
use "quotient_def.ML"
+ − 164
+ − 165
section {* Simset setup *}
+ − 166
614
+ − 167
(* Since HOL_basic_ss is too "big" for us, *)
+ − 168
(* we set up our own minimal simpset. *)
597
+ − 169
ML {*
+ − 170
fun mk_minimal_ss ctxt =
+ − 171
Simplifier.context ctxt empty_ss
+ − 172
setsubgoaler asm_simp_tac
+ − 173
setmksimps (mksimps [])
+ − 174
*}
+ − 175
605
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 176
ML {*
610
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 177
fun OF1 thm1 thm2 = thm2 RS thm1
605
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 178
*}
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 179
614
+ − 180
section {* Atomize Infrastructure *}
597
+ − 181
+ − 182
lemma atomize_eqv[atomize]:
+ − 183
shows "(Trueprop A \<equiv> Trueprop B) \<equiv> (A \<equiv> B)"
+ − 184
proof
+ − 185
assume "A \<equiv> B"
+ − 186
then show "Trueprop A \<equiv> Trueprop B" by unfold
+ − 187
next
+ − 188
assume *: "Trueprop A \<equiv> Trueprop B"
+ − 189
have "A = B"
+ − 190
proof (cases A)
+ − 191
case True
+ − 192
have "A" by fact
+ − 193
then show "A = B" using * by simp
+ − 194
next
+ − 195
case False
+ − 196
have "\<not>A" by fact
+ − 197
then show "A = B" using * by auto
+ − 198
qed
+ − 199
then show "A \<equiv> B" by (rule eq_reflection)
+ − 200
qed
+ − 201
+ − 202
ML {*
+ − 203
fun atomize_thm thm =
+ − 204
let
+ − 205
val thm' = Thm.freezeT (forall_intr_vars thm)
+ − 206
val thm'' = ObjectLogic.atomize (cprop_of thm')
+ − 207
in
+ − 208
@{thm equal_elim_rule1} OF [thm'', thm']
+ − 209
end
+ − 210
*}
+ − 211
614
+ − 212
section {* Infrastructure about id *}
+ − 213
600
+ − 214
(* TODO: think where should this be *)
597
+ − 215
lemma prod_fun_id: "prod_fun id id \<equiv> id"
+ − 216
by (rule eq_reflection) (simp add: prod_fun_def)
+ − 217
614
+ − 218
lemmas [id_simps] =
597
+ − 219
fun_map_id[THEN eq_reflection]
+ − 220
id_apply[THEN eq_reflection]
+ − 221
id_def[THEN eq_reflection,symmetric]
600
+ − 222
prod_fun_id
597
+ − 223
+ − 224
section {* Debugging infrastructure for testing tactics *}
+ − 225
+ − 226
ML {*
+ − 227
fun my_print_tac ctxt s i thm =
+ − 228
let
+ − 229
val prem_str = nth (prems_of thm) (i - 1)
+ − 230
|> Syntax.string_of_term ctxt
+ − 231
handle Subscript => "no subgoal"
+ − 232
val _ = tracing (s ^ "\n" ^ prem_str)
+ − 233
in
+ − 234
Seq.single thm
+ − 235
end *}
+ − 236
+ − 237
ML {*
+ − 238
fun DT ctxt s tac i thm =
+ − 239
let
+ − 240
val before_goal = nth (prems_of thm) (i - 1)
+ − 241
|> Syntax.string_of_term ctxt
+ − 242
val before_msg = ["before: " ^ s, before_goal, "after: " ^ s]
+ − 243
|> cat_lines
+ − 244
in
+ − 245
EVERY [tac i, my_print_tac ctxt before_msg i] thm
+ − 246
end
+ − 247
+ − 248
fun NDT ctxt s tac thm = tac thm
+ − 249
*}
+ − 250
615
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 251
section {* Matching of Terms and Types *}
597
+ − 252
+ − 253
ML {*
+ − 254
fun matches_typ (ty, ty') =
+ − 255
case (ty, ty') of
+ − 256
(_, TVar _) => true
+ − 257
| (TFree x, TFree x') => x = x'
+ − 258
| (Type (s, tys), Type (s', tys')) =>
+ − 259
s = s' andalso
+ − 260
if (length tys = length tys')
+ − 261
then (List.all matches_typ (tys ~~ tys'))
+ − 262
else false
+ − 263
| _ => false
+ − 264
+ − 265
fun matches_term (trm, trm') =
+ − 266
case (trm, trm') of
+ − 267
(_, Var _) => true
+ − 268
| (Const (s, ty), Const (s', ty')) => s = s' andalso matches_typ (ty, ty')
+ − 269
| (Free (x, ty), Free (x', ty')) => x = x' andalso matches_typ (ty, ty')
+ − 270
| (Bound i, Bound j) => i = j
+ − 271
| (Abs (_, T, t), Abs (_, T', t')) => matches_typ (T, T') andalso matches_term (t, t')
+ − 272
| (t $ s, t' $ s') => matches_term (t, t') andalso matches_term (s, s')
+ − 273
| _ => false
+ − 274
*}
+ − 275
612
+ − 276
section {* Computation of the Regularize Goal *}
597
+ − 277
+ − 278
(*
+ − 279
Regularizing an rtrm means:
+ − 280
- quantifiers over a type that needs lifting are replaced by
+ − 281
bounded quantifiers, for example:
+ − 282
\<forall>x. P \<Longrightarrow> \<forall>x \<in> (Respects R). P / All (Respects R) P
+ − 283
+ − 284
the relation R is given by the rty and qty;
+ − 285
+ − 286
- abstractions over a type that needs lifting are replaced
+ − 287
by bounded abstractions:
+ − 288
\<lambda>x. P \<Longrightarrow> Ball (Respects R) (\<lambda>x. P)
+ − 289
+ − 290
- equalities over the type being lifted are replaced by
+ − 291
corresponding relations:
+ − 292
A = B \<Longrightarrow> A \<approx> B
+ − 293
+ − 294
example with more complicated types of A, B:
+ − 295
A = B \<Longrightarrow> (op = \<Longrightarrow> op \<approx>) A B
+ − 296
*)
+ − 297
+ − 298
ML {*
+ − 299
(* builds the relation that is the argument of respects *)
+ − 300
fun mk_resp_arg lthy (rty, qty) =
+ − 301
let
+ − 302
val thy = ProofContext.theory_of lthy
+ − 303
in
+ − 304
if rty = qty
+ − 305
then HOLogic.eq_const rty
+ − 306
else
+ − 307
case (rty, qty) of
+ − 308
(Type (s, tys), Type (s', tys')) =>
+ − 309
if s = s'
+ − 310
then let
+ − 311
val SOME map_info = maps_lookup thy s
+ − 312
val args = map (mk_resp_arg lthy) (tys ~~ tys')
+ − 313
in
+ − 314
list_comb (Const (#relfun map_info, dummyT), args)
+ − 315
end
+ − 316
else let
+ − 317
val SOME qinfo = quotdata_lookup_thy thy s'
+ − 318
(* FIXME: check in this case that the rty and qty *)
+ − 319
(* FIXME: correspond to each other *)
+ − 320
val (s, _) = dest_Const (#rel qinfo)
+ − 321
(* FIXME: the relation should only be the string *)
+ − 322
(* FIXME: and the type needs to be calculated as below; *)
+ − 323
(* FIXME: maybe one should actually have a term *)
+ − 324
(* FIXME: and one needs to force it to have this type *)
+ − 325
in
+ − 326
Const (s, rty --> rty --> @{typ bool})
+ − 327
end
+ − 328
| _ => HOLogic.eq_const dummyT
+ − 329
(* FIXME: check that the types correspond to each other? *)
+ − 330
end
+ − 331
*}
+ − 332
+ − 333
ML {*
+ − 334
val mk_babs = Const (@{const_name Babs}, dummyT)
+ − 335
val mk_ball = Const (@{const_name Ball}, dummyT)
+ − 336
val mk_bex = Const (@{const_name Bex}, dummyT)
+ − 337
val mk_resp = Const (@{const_name Respects}, dummyT)
+ − 338
*}
+ − 339
+ − 340
ML {*
+ − 341
(* - applies f to the subterm of an abstraction, *)
+ − 342
(* otherwise to the given term, *)
+ − 343
(* - used by regularize, therefore abstracted *)
+ − 344
(* variables do not have to be treated specially *)
+ − 345
+ − 346
fun apply_subt f trm1 trm2 =
+ − 347
case (trm1, trm2) of
+ − 348
(Abs (x, T, t), Abs (x', T', t')) => Abs (x, T, f t t')
+ − 349
| _ => f trm1 trm2
+ − 350
+ − 351
(* the major type of All and Ex quantifiers *)
+ − 352
fun qnt_typ ty = domain_type (domain_type ty)
+ − 353
*}
+ − 354
+ − 355
ML {*
632
+ − 356
(* produces a regularized version of rtrm *)
+ − 357
(* - the result is contains dummyT *)
597
+ − 358
(* - does not need any special treatment of *)
+ − 359
(* bound variables *)
+ − 360
+ − 361
fun regularize_trm lthy rtrm qtrm =
+ − 362
case (rtrm, qtrm) of
+ − 363
(Abs (x, ty, t), Abs (x', ty', t')) =>
+ − 364
let
+ − 365
val subtrm = Abs(x, ty, regularize_trm lthy t t')
+ − 366
in
632
+ − 367
if ty = ty' then subtrm
597
+ − 368
else mk_babs $ (mk_resp $ mk_resp_arg lthy (ty, ty')) $ subtrm
+ − 369
end
+ − 370
+ − 371
| (Const (@{const_name "All"}, ty) $ t, Const (@{const_name "All"}, ty') $ t') =>
+ − 372
let
+ − 373
val subtrm = apply_subt (regularize_trm lthy) t t'
+ − 374
in
655
+ − 375
if ty = ty' then Const (@{const_name "All"}, ty) $ subtrm
597
+ − 376
else mk_ball $ (mk_resp $ mk_resp_arg lthy (qnt_typ ty, qnt_typ ty')) $ subtrm
+ − 377
end
+ − 378
+ − 379
| (Const (@{const_name "Ex"}, ty) $ t, Const (@{const_name "Ex"}, ty') $ t') =>
+ − 380
let
+ − 381
val subtrm = apply_subt (regularize_trm lthy) t t'
+ − 382
in
655
+ − 383
if ty = ty' then Const (@{const_name "Ex"}, ty) $ subtrm
597
+ − 384
else mk_bex $ (mk_resp $ mk_resp_arg lthy (qnt_typ ty, qnt_typ ty')) $ subtrm
+ − 385
end
+ − 386
+ − 387
| (* equalities need to be replaced by appropriate equivalence relations *)
+ − 388
(Const (@{const_name "op ="}, ty), Const (@{const_name "op ="}, ty')) =>
632
+ − 389
if ty = ty' then rtrm
597
+ − 390
else mk_resp_arg lthy (domain_type ty, domain_type ty')
+ − 391
+ − 392
| (* in this case we check whether the given equivalence relation is correct *)
+ − 393
(rel, Const (@{const_name "op ="}, ty')) =>
+ − 394
let
+ − 395
val exc = LIFT_MATCH "regularise (relation mismatch)"
+ − 396
val rel_ty = (fastype_of rel) handle TERM _ => raise exc
+ − 397
val rel' = mk_resp_arg lthy (domain_type rel_ty, domain_type ty')
+ − 398
in
655
+ − 399
if rel' = rel then rtrm else raise exc
597
+ − 400
end
+ − 401
| (_, Const (s, _)) =>
+ − 402
let
+ − 403
fun same_name (Const (s, _)) (Const (s', _)) = (s = s')
+ − 404
| same_name _ _ = false
+ − 405
in
+ − 406
if same_name rtrm qtrm
+ − 407
then rtrm
+ − 408
else
+ − 409
let
+ − 410
fun exc1 s = LIFT_MATCH ("regularize (constant " ^ s ^ " not found)")
+ − 411
val exc2 = LIFT_MATCH ("regularize (constant mismatch)")
+ − 412
val thy = ProofContext.theory_of lthy
+ − 413
val rtrm' = (#rconst (qconsts_lookup thy s)) handle NotFound => raise (exc1 s)
+ − 414
in
655
+ − 415
if matches_term (rtrm, rtrm') then rtrm else raise exc2
597
+ − 416
end
+ − 417
end
+ − 418
+ − 419
| (t1 $ t2, t1' $ t2') =>
+ − 420
(regularize_trm lthy t1 t1') $ (regularize_trm lthy t2 t2')
+ − 421
+ − 422
| (Free (x, ty), Free (x', ty')) =>
+ − 423
(* this case cannot arrise as we start with two fully atomized terms *)
+ − 424
raise (LIFT_MATCH "regularize (frees)")
+ − 425
+ − 426
| (Bound i, Bound i') =>
632
+ − 427
if i = i' then rtrm
597
+ − 428
else raise (LIFT_MATCH "regularize (bounds mismatch)")
+ − 429
+ − 430
| (rt, qt) =>
632
+ − 431
raise (LIFT_MATCH "regularize failed (default)")
597
+ − 432
*}
+ − 433
612
+ − 434
section {* Regularize Tactic *}
+ − 435
597
+ − 436
ML {*
+ − 437
fun equiv_tac ctxt =
605
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 438
REPEAT_ALL_NEW (resolve_tac (equiv_rules_get ctxt))
597
+ − 439
+ − 440
fun equiv_solver_tac ss = equiv_tac (Simplifier.the_context ss)
+ − 441
val equiv_solver = Simplifier.mk_solver' "Equivalence goal solver" equiv_solver_tac
+ − 442
*}
+ − 443
+ − 444
ML {*
+ − 445
fun prep_trm thy (x, (T, t)) =
+ − 446
(cterm_of thy (Var (x, T)), cterm_of thy t)
+ − 447
+ − 448
fun prep_ty thy (x, (S, ty)) =
+ − 449
(ctyp_of thy (TVar (x, S)), ctyp_of thy ty)
+ − 450
*}
+ − 451
+ − 452
ML {*
+ − 453
fun matching_prs thy pat trm =
+ − 454
let
+ − 455
val univ = Unify.matchers thy [(pat, trm)]
+ − 456
val SOME (env, _) = Seq.pull univ
+ − 457
val tenv = Vartab.dest (Envir.term_env env)
+ − 458
val tyenv = Vartab.dest (Envir.type_env env)
+ − 459
in
+ − 460
(map (prep_ty thy) tyenv, map (prep_trm thy) tenv)
+ − 461
end
+ − 462
*}
+ − 463
+ − 464
ML {*
+ − 465
fun calculate_instance ctxt thm redex R1 R2 =
+ − 466
let
+ − 467
val thy = ProofContext.theory_of ctxt
+ − 468
val goal = Const (@{const_name "equivp"}, dummyT) $ R2
+ − 469
|> Syntax.check_term ctxt
+ − 470
|> HOLogic.mk_Trueprop
+ − 471
val eqv_prem = Goal.prove ctxt [] [] goal (fn {context,...} => equiv_tac context 1)
+ − 472
val thm = (@{thm eq_reflection} OF [thm OF [eqv_prem]])
+ − 473
val R1c = cterm_of thy R1
+ − 474
val thmi = Drule.instantiate' [] [SOME R1c] thm
+ − 475
val inst = matching_prs thy (term_of (Thm.lhs_of thmi)) redex
+ − 476
val thm2 = Drule.eta_contraction_rule (Drule.instantiate inst thmi)
+ − 477
in
+ − 478
SOME thm2
+ − 479
end
+ − 480
handle _ => NONE
615
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 481
(* FIXME/TODO: what is the place where the exception is raised: matching_prs? *)
597
+ − 482
*}
+ − 483
+ − 484
ML {*
+ − 485
fun ball_bex_range_simproc ss redex =
+ − 486
let
+ − 487
val ctxt = Simplifier.the_context ss
+ − 488
in
+ − 489
case redex of
+ − 490
(Const (@{const_name "Ball"}, _) $ (Const (@{const_name "Respects"}, _) $
+ − 491
(Const (@{const_name "fun_rel"}, _) $ R1 $ R2)) $ _) =>
+ − 492
calculate_instance ctxt @{thm ball_reg_eqv_range} redex R1 R2
+ − 493
| (Const (@{const_name "Bex"}, _) $ (Const (@{const_name "Respects"}, _) $
+ − 494
(Const (@{const_name "fun_rel"}, _) $ R1 $ R2)) $ _) =>
+ − 495
calculate_instance ctxt @{thm bex_reg_eqv_range} redex R1 R2
+ − 496
| _ => NONE
+ − 497
end
+ − 498
*}
+ − 499
+ − 500
lemma eq_imp_rel:
+ − 501
shows "equivp R \<Longrightarrow> a = b \<longrightarrow> R a b"
+ − 502
by (simp add: equivp_reflp)
+ − 503
+ − 504
612
+ − 505
(* Regularize Tactic *)
597
+ − 506
612
+ − 507
(* 0. preliminary simplification step according to *)
+ − 508
thm ball_reg_eqv bex_reg_eqv babs_reg_eqv (* the latter of no use *)
+ − 509
ball_reg_eqv_range bex_reg_eqv_range
+ − 510
(* 1. eliminating simple Ball/Bex instances*)
+ − 511
thm ball_reg_right bex_reg_left
+ − 512
(* 2. monos *)
+ − 513
(* 3. commutation rules for ball and bex *)
+ − 514
thm ball_all_comm bex_ex_comm
+ − 515
(* 4. then rel-equality (which need to be instantiated to avoid loops *)
+ − 516
thm eq_imp_rel
+ − 517
(* 5. then simplification like 0 *)
+ − 518
(* finally jump back to 1 *)
597
+ − 519
612
+ − 520
597
+ − 521
ML {*
+ − 522
fun regularize_tac ctxt =
+ − 523
let
+ − 524
val thy = ProofContext.theory_of ctxt
+ − 525
val pat_ball = @{term "Ball (Respects (R1 ===> R2)) P"}
+ − 526
val pat_bex = @{term "Bex (Respects (R1 ===> R2)) P"}
+ − 527
val simproc = Simplifier.simproc_i thy "" [pat_ball, pat_bex] (K (ball_bex_range_simproc))
+ − 528
val simpset = (mk_minimal_ss ctxt)
605
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 529
addsimps @{thms ball_reg_eqv bex_reg_eqv babs_reg_eqv}
597
+ − 530
addsimprocs [simproc] addSolver equiv_solver
+ − 531
(* TODO: Make sure that there are no list_rel, pair_rel etc involved *)
605
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 532
(* can this cause loops in equiv_tac ? *)
612
+ − 533
val eq_eqvs = map (OF1 @{thm eq_imp_rel}) (equiv_rules_get ctxt)
597
+ − 534
in
+ − 535
simp_tac simpset THEN'
605
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 536
REPEAT_ALL_NEW (CHANGED o FIRST' [
612
+ − 537
resolve_tac @{thms ball_reg_right bex_reg_left},
597
+ − 538
resolve_tac (Inductive.get_monos ctxt),
612
+ − 539
resolve_tac @{thms ball_all_comm bex_ex_comm},
605
120e479ed367
first attempt to deal with Babs in regularise and cleaning (not yet working)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 540
resolve_tac eq_eqvs,
597
+ − 541
simp_tac simpset])
+ − 542
end
+ − 543
*}
+ − 544
612
+ − 545
section {* Calculation of the Injected Goal *}
597
+ − 546
+ − 547
(*
+ − 548
Injecting repabs means:
+ − 549
+ − 550
For abstractions:
+ − 551
* If the type of the abstraction doesn't need lifting we recurse.
+ − 552
* If it does we add RepAbs around the whole term and check if the
+ − 553
variable needs lifting.
+ − 554
* If it doesn't then we recurse
+ − 555
* If it does we recurse and put 'RepAbs' around all occurences
+ − 556
of the variable in the obtained subterm. This in combination
+ − 557
with the RepAbs above will let us change the type of the
+ − 558
abstraction with rewriting.
+ − 559
For applications:
+ − 560
* If the term is 'Respects' applied to anything we leave it unchanged
+ − 561
* If the term needs lifting and the head is a constant that we know
+ − 562
how to lift, we put a RepAbs and recurse
+ − 563
* If the term needs lifting and the head is a free applied to subterms
+ − 564
(if it is not applied we treated it in Abs branch) then we
+ − 565
put RepAbs and recurse
+ − 566
* Otherwise just recurse.
+ − 567
*)
+ − 568
+ − 569
ML {*
+ − 570
fun mk_repabs lthy (T, T') trm =
+ − 571
Quotient_Def.get_fun repF lthy (T, T')
+ − 572
$ (Quotient_Def.get_fun absF lthy (T, T') $ trm)
+ − 573
*}
+ − 574
+ − 575
ML {*
+ − 576
(* bound variables need to be treated properly, *)
+ − 577
(* as the type of subterms need to be calculated *)
+ − 578
(* in the abstraction case *)
+ − 579
+ − 580
fun inj_repabs_trm lthy (rtrm, qtrm) =
+ − 581
case (rtrm, qtrm) of
+ − 582
(Const (@{const_name "Ball"}, T) $ r $ t, Const (@{const_name "All"}, _) $ t') =>
+ − 583
Const (@{const_name "Ball"}, T) $ r $ (inj_repabs_trm lthy (t, t'))
+ − 584
+ − 585
| (Const (@{const_name "Bex"}, T) $ r $ t, Const (@{const_name "Ex"}, _) $ t') =>
+ − 586
Const (@{const_name "Bex"}, T) $ r $ (inj_repabs_trm lthy (t, t'))
+ − 587
+ − 588
| (Const (@{const_name "Babs"}, T) $ r $ t, t' as (Abs _)) =>
621
c10a46fa0de9
Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 589
let
c10a46fa0de9
Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 590
val rty = fastype_of rtrm
c10a46fa0de9
Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 591
val qty = fastype_of qtrm
c10a46fa0de9
Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 592
in
c10a46fa0de9
Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 593
mk_repabs lthy (rty, qty) (Const (@{const_name "Babs"}, T) $ r $ (inj_repabs_trm lthy (t, t')))
c10a46fa0de9
Added a 'rep_abs' in inj_repabs_trm of babs; and proved two lam examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 594
end
597
+ − 595
+ − 596
| (Abs (x, T, t), Abs (x', T', t')) =>
+ − 597
let
+ − 598
val rty = fastype_of rtrm
+ − 599
val qty = fastype_of qtrm
+ − 600
val (y, s) = Term.dest_abs (x, T, t)
+ − 601
val (_, s') = Term.dest_abs (x', T', t')
+ − 602
val yvar = Free (y, T)
+ − 603
val result = Term.lambda_name (y, yvar) (inj_repabs_trm lthy (s, s'))
+ − 604
in
615
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 605
if rty = qty then result
597
+ − 606
else mk_repabs lthy (rty, qty) result
+ − 607
end
+ − 608
+ − 609
| (t $ s, t' $ s') =>
+ − 610
(inj_repabs_trm lthy (t, t')) $ (inj_repabs_trm lthy (s, s'))
+ − 611
+ − 612
| (Free (_, T), Free (_, T')) =>
615
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 613
if T = T' then rtrm
597
+ − 614
else mk_repabs lthy (T, T') rtrm
+ − 615
+ − 616
| (_, Const (@{const_name "op ="}, _)) => rtrm
+ − 617
+ − 618
(* FIXME: check here that rtrm is the corresponding definition for the const *)
+ − 619
| (_, Const (_, T')) =>
+ − 620
let
+ − 621
val rty = fastype_of rtrm
+ − 622
in
615
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 623
if rty = T' then rtrm
597
+ − 624
else mk_repabs lthy (rty, T') rtrm
+ − 625
end
+ − 626
+ − 627
| _ => raise (LIFT_MATCH "injection")
+ − 628
*}
+ − 629
612
+ − 630
section {* Injection Tactic *}
597
+ − 631
+ − 632
ML {*
+ − 633
fun quotient_tac ctxt =
+ − 634
REPEAT_ALL_NEW (FIRST'
+ − 635
[rtac @{thm identity_quotient},
+ − 636
resolve_tac (quotient_rules_get ctxt)])
+ − 637
+ − 638
fun quotient_solver_tac ss = quotient_tac (Simplifier.the_context ss)
+ − 639
val quotient_solver = Simplifier.mk_solver' "Quotient goal solver" quotient_solver_tac
+ − 640
*}
+ − 641
+ − 642
ML {*
+ − 643
fun solve_quotient_assums ctxt thm =
610
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 644
let
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 645
val goal = hd (Drule.strip_imp_prems (cprop_of thm))
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 646
in
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 647
thm OF [Goal.prove_internal [] goal (fn _ => quotient_tac ctxt 1)]
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 648
end
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 649
handle _ => error "solve_quotient_assums failed. Maybe a quotient_thm is missing"
597
+ − 650
*}
+ − 651
+ − 652
(* Not used *)
+ − 653
(* It proves the Quotient assumptions by calling quotient_tac *)
+ − 654
ML {*
+ − 655
fun solve_quotient_assum i ctxt thm =
+ − 656
let
+ − 657
val tac =
+ − 658
(compose_tac (false, thm, i)) THEN_ALL_NEW
+ − 659
(quotient_tac ctxt);
+ − 660
val gc = Drule.strip_imp_concl (cprop_of thm);
+ − 661
in
+ − 662
Goal.prove_internal [] gc (fn _ => tac 1)
+ − 663
end
+ − 664
handle _ => error "solve_quotient_assum"
+ − 665
*}
+ − 666
+ − 667
definition
+ − 668
"QUOT_TRUE x \<equiv> True"
+ − 669
+ − 670
ML {*
+ − 671
fun find_qt_asm asms =
+ − 672
let
+ − 673
fun find_fun trm =
+ − 674
case trm of
+ − 675
(Const(@{const_name Trueprop}, _) $ (Const (@{const_name QUOT_TRUE}, _) $ _)) => true
+ − 676
| _ => false
+ − 677
in
+ − 678
case find_first find_fun asms of
+ − 679
SOME (_ $ (_ $ (f $ a))) => (f, a)
+ − 680
| SOME _ => error "find_qt_asm: no pair"
+ − 681
| NONE => error "find_qt_asm: no assumption"
+ − 682
end
+ − 683
*}
+ − 684
+ − 685
(*
+ − 686
To prove that the regularised theorem implies the abs/rep injected,
+ − 687
we try:
+ − 688
+ − 689
1) theorems 'trans2' from the appropriate QUOT_TYPE
+ − 690
2) remove lambdas from both sides: lambda_rsp_tac
+ − 691
3) remove Ball/Bex from the right hand side
+ − 692
4) use user-supplied RSP theorems
+ − 693
5) remove rep_abs from the right side
+ − 694
6) reflexivity of equality
+ − 695
7) split applications of lifted type (apply_rsp)
+ − 696
8) split applications of non-lifted type (cong_tac)
+ − 697
9) apply extentionality
+ − 698
A) reflexivity of the relation
+ − 699
B) assumption
+ − 700
(Lambdas under respects may have left us some assumptions)
+ − 701
C) proving obvious higher order equalities by simplifying fun_rel
+ − 702
(not sure if it is still needed?)
+ − 703
D) unfolding lambda on one side
+ − 704
E) simplifying (= ===> =) for simpler respectfulness
+ − 705
+ − 706
*)
+ − 707
+ − 708
lemma quot_true_dests:
+ − 709
shows QT_all: "QUOT_TRUE (All P) \<Longrightarrow> QUOT_TRUE P"
+ − 710
and QT_ex: "QUOT_TRUE (Ex P) \<Longrightarrow> QUOT_TRUE P"
+ − 711
and QT_lam: "QUOT_TRUE (\<lambda>x. P x) \<Longrightarrow> (\<And>x. QUOT_TRUE (P x))"
+ − 712
and QT_ext: "(\<And>x. QUOT_TRUE (a x) \<Longrightarrow> f x = g x) \<Longrightarrow> (QUOT_TRUE a \<Longrightarrow> f = g)"
655
+ − 713
by (simp_all add: QUOT_TRUE_def ext)
597
+ − 714
+ − 715
lemma QUOT_TRUE_i: "(QUOT_TRUE (a :: bool) \<Longrightarrow> P) \<Longrightarrow> P"
+ − 716
by (simp add: QUOT_TRUE_def)
+ − 717
+ − 718
lemma QUOT_TRUE_imp: "QUOT_TRUE a \<equiv> QUOT_TRUE b"
+ − 719
by (simp add: QUOT_TRUE_def)
+ − 720
+ − 721
ML {*
+ − 722
fun quot_true_conv1 ctxt fnctn ctrm =
+ − 723
case (term_of ctrm) of
+ − 724
(Const (@{const_name QUOT_TRUE}, _) $ x) =>
+ − 725
let
+ − 726
val fx = fnctn x;
+ − 727
val thy = ProofContext.theory_of ctxt;
+ − 728
val cx = cterm_of thy x;
+ − 729
val cfx = cterm_of thy fx;
+ − 730
val cxt = ctyp_of thy (fastype_of x);
+ − 731
val cfxt = ctyp_of thy (fastype_of fx);
+ − 732
val thm = Drule.instantiate' [SOME cxt, SOME cfxt] [SOME cx, SOME cfx] @{thm QUOT_TRUE_imp}
+ − 733
in
+ − 734
Conv.rewr_conv thm ctrm
+ − 735
end
+ − 736
*}
+ − 737
+ − 738
ML {*
+ − 739
fun quot_true_conv ctxt fnctn ctrm =
+ − 740
case (term_of ctrm) of
+ − 741
(Const (@{const_name QUOT_TRUE}, _) $ _) =>
+ − 742
quot_true_conv1 ctxt fnctn ctrm
+ − 743
| _ $ _ => Conv.comb_conv (quot_true_conv ctxt fnctn) ctrm
+ − 744
| Abs _ => Conv.abs_conv (fn (_, ctxt) => quot_true_conv ctxt fnctn) ctxt ctrm
+ − 745
| _ => Conv.all_conv ctrm
+ − 746
*}
+ − 747
+ − 748
ML {*
+ − 749
fun quot_true_tac ctxt fnctn = CONVERSION
+ − 750
((Conv.params_conv ~1 (fn ctxt =>
+ − 751
(Conv.prems_conv ~1 (quot_true_conv ctxt fnctn)))) ctxt)
+ − 752
*}
+ − 753
+ − 754
ML {* fun dest_comb (f $ a) = (f, a) *}
+ − 755
ML {* fun dest_bcomb ((_ $ l) $ r) = (l, r) *}
+ − 756
(* TODO: Can this be done easier? *)
+ − 757
ML {*
+ − 758
fun unlam t =
+ − 759
case t of
+ − 760
(Abs a) => snd (Term.dest_abs a)
+ − 761
| _ => unlam (Abs("", domain_type (fastype_of t), (incr_boundvars 1 t) $ (Bound 0)))
+ − 762
*}
+ − 763
+ − 764
ML {*
+ − 765
fun dest_fun_type (Type("fun", [T, S])) = (T, S)
+ − 766
| dest_fun_type _ = error "dest_fun_type"
+ − 767
*}
+ − 768
+ − 769
ML {*
+ − 770
val bare_concl = HOLogic.dest_Trueprop o Logic.strip_assums_concl
+ − 771
*}
+ − 772
+ − 773
ML {*
+ − 774
val apply_rsp_tac =
+ − 775
Subgoal.FOCUS (fn {concl, asms, context,...} =>
+ − 776
case ((HOLogic.dest_Trueprop (term_of concl))) of
+ − 777
((R2 $ (f $ x) $ (g $ y))) =>
+ − 778
(let
+ − 779
val (asmf, asma) = find_qt_asm (map term_of asms);
+ − 780
in
+ − 781
if (fastype_of asmf) = (fastype_of f) then no_tac else let
+ − 782
val ty_a = fastype_of x;
+ − 783
val ty_b = fastype_of asma;
+ − 784
val ty_c = range_type (type_of f);
+ − 785
val thy = ProofContext.theory_of context;
+ − 786
val ty_inst = map (SOME o (ctyp_of thy)) [ty_a, ty_b, ty_c];
+ − 787
val thm = Drule.instantiate' ty_inst [] @{thm apply_rsp}
+ − 788
val te = solve_quotient_assums context thm
+ − 789
val t_inst = map (SOME o (cterm_of thy)) [R2, f, g, x, y];
+ − 790
val thm = Drule.instantiate' [] t_inst te
+ − 791
in
+ − 792
compose_tac (false, thm, 2) 1
+ − 793
end
+ − 794
end
+ − 795
handle ERROR "find_qt_asm: no pair" => no_tac)
+ − 796
| _ => no_tac)
+ − 797
*}
+ − 798
+ − 799
ML {*
629
+ − 800
fun equals_rsp_tac R ctxt =
+ − 801
let
+ − 802
val t = domain_type (fastype_of R);
+ − 803
val thy = ProofContext.theory_of ctxt
+ − 804
val thm = Drule.instantiate' [SOME (ctyp_of thy t)] [SOME (cterm_of thy R)] @{thm equals_rsp}
+ − 805
in
+ − 806
rtac thm THEN' RANGE [quotient_tac ctxt]
+ − 807
end
+ − 808
handle THM _ => K no_tac | TYPE _ => K no_tac | TERM _ => K no_tac
+ − 809
*}
+ − 810
+ − 811
ML {*
597
+ − 812
fun rep_abs_rsp_tac ctxt =
+ − 813
SUBGOAL (fn (goal, i) =>
+ − 814
case (bare_concl goal) of
+ − 815
(rel $ _ $ (rep $ (abs $ _))) =>
+ − 816
(let
+ − 817
val thy = ProofContext.theory_of ctxt;
+ − 818
val (ty_a, ty_b) = dest_fun_type (fastype_of abs);
+ − 819
val ty_inst = map (SOME o (ctyp_of thy)) [ty_a, ty_b];
+ − 820
val t_inst = map (SOME o (cterm_of thy)) [rel, abs, rep];
+ − 821
val thm = Drule.instantiate' ty_inst t_inst @{thm rep_abs_rsp}
+ − 822
val te = solve_quotient_assums ctxt thm
+ − 823
in
+ − 824
rtac te i
+ − 825
end
+ − 826
handle _ => no_tac)
+ − 827
| _ => no_tac)
+ − 828
*}
+ − 829
+ − 830
ML {*
629
+ − 831
fun inj_repabs_tac_match ctxt = SUBGOAL (fn (goal, i) =>
597
+ − 832
(case (bare_concl goal) of
+ − 833
(* (R1 ===> R2) (\<lambda>x\<dots>) (\<lambda>y\<dots>) ----> \<lbrakk>R1 x y\<rbrakk> \<Longrightarrow> R2 (\<dots>x) (\<dots>y) *)
+ − 834
((Const (@{const_name fun_rel}, _) $ _ $ _) $ (Abs _) $ (Abs _))
+ − 835
=> rtac @{thm fun_rel_id} THEN' quot_true_tac ctxt unlam
+ − 836
+ − 837
(* (op =) (Ball\<dots>) (Ball\<dots>) ----> (op =) (\<dots>) (\<dots>) *)
+ − 838
| (Const (@{const_name "op ="},_) $
+ − 839
(Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
+ − 840
(Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _))
+ − 841
=> rtac @{thm ball_rsp} THEN' dtac @{thm QT_all}
+ − 842
+ − 843
(* (R1 ===> op =) (Ball\<dots>) (Ball\<dots>) ----> \<lbrakk>R1 x y\<rbrakk> \<Longrightarrow> (Ball\<dots>x) = (Ball\<dots>y) *)
+ − 844
| (Const (@{const_name fun_rel}, _) $ _ $ _) $
+ − 845
(Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
+ − 846
(Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _)
+ − 847
=> rtac @{thm fun_rel_id} THEN' quot_true_tac ctxt unlam
+ − 848
+ − 849
(* (op =) (Bex\<dots>) (Bex\<dots>) ----> (op =) (\<dots>) (\<dots>) *)
+ − 850
| Const (@{const_name "op ="},_) $
+ − 851
(Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
+ − 852
(Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _)
+ − 853
=> rtac @{thm bex_rsp} THEN' dtac @{thm QT_ex}
+ − 854
+ − 855
(* (R1 ===> op =) (Bex\<dots>) (Bex\<dots>) ----> \<lbrakk>R1 x y\<rbrakk> \<Longrightarrow> (Bex\<dots>x) = (Bex\<dots>y) *)
+ − 856
| (Const (@{const_name fun_rel}, _) $ _ $ _) $
+ − 857
(Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
+ − 858
(Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _)
+ − 859
=> rtac @{thm fun_rel_id} THEN' quot_true_tac ctxt unlam
+ − 860
+ − 861
| (_ $
+ − 862
(Const(@{const_name Babs},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
+ − 863
(Const(@{const_name Babs},_) $ (Const (@{const_name Respects}, _) $ _) $ _))
+ − 864
=> rtac @{thm babs_rsp} THEN' RANGE [quotient_tac ctxt]
+ − 865
629
+ − 866
| Const (@{const_name "op ="},_) $ (R $ _ $ _) $ (_ $ _ $ _) => (rtac @{thm refl} ORELSE'
+ − 867
(equals_rsp_tac R ctxt THEN' RANGE [
+ − 868
quot_true_tac ctxt (fst o dest_bcomb), quot_true_tac ctxt (snd o dest_bcomb)]))
+ − 869
597
+ − 870
(* reflexivity of operators arising from Cong_tac *)
629
+ − 871
| Const (@{const_name "op ="},_) $ _ $ _ => rtac @{thm refl}
597
+ − 872
+ − 873
(* respectfulness of constants; in particular of a simple relation *)
+ − 874
| _ $ (Const _) $ (Const _) (* fun_rel, list_rel, etc but not equality *)
+ − 875
=> resolve_tac (rsp_rules_get ctxt) THEN_ALL_NEW quotient_tac ctxt
+ − 876
+ − 877
(* R (\<dots>) (Rep (Abs \<dots>)) ----> R (\<dots>) (\<dots>) *)
+ − 878
(* observe ---> *)
624
c4299ce27e46
Removed pattern from quot_rel_rsp, since list_rel and all used introduced ones cannot be patterned
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 879
| _ $ _ $ _
c4299ce27e46
Removed pattern from quot_rel_rsp, since list_rel and all used introduced ones cannot be patterned
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 880
=> (rtac @{thm quot_rel_rsp} THEN_ALL_NEW quotient_tac ctxt) ORELSE' rep_abs_rsp_tac ctxt
597
+ − 881
+ − 882
| _ => error "inj_repabs_tac not a relation"
+ − 883
) i)
+ − 884
*}
+ − 885
+ − 886
ML {*
629
+ − 887
fun inj_repabs_step_tac ctxt rel_refl =
597
+ − 888
(FIRST' [
629
+ − 889
NDT ctxt "0" (inj_repabs_tac_match ctxt),
597
+ − 890
(* R (t $ \<dots>) (t' $ \<dots>) ----> apply_rsp provided type of t needs lifting *)
610
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 891
597
+ − 892
NDT ctxt "A" (apply_rsp_tac ctxt THEN'
610
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 893
RANGE [quot_true_tac ctxt (fst o dest_comb), quot_true_tac ctxt (snd o dest_comb)]),
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 894
597
+ − 895
(* (op =) (t $ \<dots>) (t' $ \<dots>) ----> Cong provided type of t does not need lifting *)
+ − 896
(* merge with previous tactic *)
610
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 897
NDT ctxt "C" (Cong_Tac.cong_tac @{thm cong} THEN'
597
+ − 898
(RANGE [quot_true_tac ctxt (fst o dest_comb), quot_true_tac ctxt (snd o dest_comb)])),
610
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 899
597
+ − 900
(* (op =) (\<lambda>x\<dots>) (\<lambda>x\<dots>) ----> (op =) (\<dots>) (\<dots>) *)
610
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 901
NDT ctxt "D" (rtac @{thm ext} THEN' quot_true_tac ctxt unlam),
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 902
597
+ − 903
(* resolving with R x y assumptions *)
+ − 904
NDT ctxt "E" (atac),
610
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 905
597
+ − 906
(* reflexivity of the basic relations *)
+ − 907
(* R \<dots> \<dots> *)
610
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 908
NDT ctxt "F" (resolve_tac rel_refl)
597
+ − 909
])
+ − 910
*}
+ − 911
+ − 912
ML {*
610
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 913
fun inj_repabs_tac ctxt =
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 914
let
612
+ − 915
val rel_refl = map (OF1 @{thm equivp_reflp}) (equiv_rules_get ctxt)
610
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 916
in
629
+ − 917
inj_repabs_step_tac ctxt rel_refl
610
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 918
end
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 919
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 920
fun all_inj_repabs_tac ctxt =
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 921
REPEAT_ALL_NEW (inj_repabs_tac ctxt)
597
+ − 922
*}
+ − 923
615
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 924
section {* Cleaning of the Theorem *}
597
+ − 925
659
+ − 926
ML {*
+ − 927
fun fun_map_simple_conv xs ctxt ctrm =
+ − 928
case (term_of ctrm) of
+ − 929
((Const (@{const_name "fun_map"}, _) $ _ $ _) $ h $ _) =>
+ − 930
if (member (op=) xs h)
+ − 931
then Conv.all_conv ctrm
+ − 932
else Conv.rewr_conv @{thm fun_map.simps[THEN eq_reflection]} ctrm
+ − 933
| _ => Conv.all_conv ctrm
+ − 934
+ − 935
fun fun_map_conv xs ctxt ctrm =
+ − 936
case (term_of ctrm) of
+ − 937
_ $ _ => (Conv.comb_conv (fun_map_conv xs ctxt) then_conv
+ − 938
fun_map_simple_conv xs ctxt) ctrm
+ − 939
| Abs _ => Conv.abs_conv (fn (x, ctxt) => fun_map_conv ((term_of x)::xs) ctxt) ctxt ctrm
+ − 940
| _ => Conv.all_conv ctrm
+ − 941
+ − 942
fun fun_map_tac ctxt = CONVERSION (fun_map_conv [] ctxt)
+ − 943
*}
+ − 944
631
+ − 945
(* Since the patterns for the lhs are different; there are 3 different make-insts *)
+ − 946
(* 1: does ? \<rightarrow> id *)
+ − 947
(* 2: does id \<rightarrow> ? *)
+ − 948
(* 3: does ? \<rightarrow> ? *)
597
+ − 949
ML {*
+ − 950
fun make_inst lhs t =
+ − 951
let
+ − 952
val _ $ (Abs (_, _, (f as Var (_, Type ("fun", [T, _]))) $ u)) = lhs;
+ − 953
val _ $ (Abs (_, _, g)) = t;
+ − 954
fun mk_abs i t =
+ − 955
if incr_boundvars i u aconv t then Bound i
+ − 956
else (case t of
+ − 957
t1 $ t2 => mk_abs i t1 $ mk_abs i t2
+ − 958
| Abs (s, T, t') => Abs (s, T, mk_abs (i + 1) t')
+ − 959
| Bound j => if i = j then error "make_inst" else t
+ − 960
| _ => t);
+ − 961
in (f, Abs ("x", T, mk_abs 0 g)) end;
+ − 962
*}
+ − 963
602
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 964
ML {*
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 965
fun make_inst2 lhs t =
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 966
let
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 967
val _ $ (Abs (_, _, (_ $ ((f as Var (_, Type ("fun", [T, _]))) $ u)))) = lhs;
631
+ − 968
val _ $ (Abs (_, _, (_ $ g))) = t;
+ − 969
fun mk_abs i t =
+ − 970
if incr_boundvars i u aconv t then Bound i
+ − 971
else (case t of
+ − 972
t1 $ t2 => mk_abs i t1 $ mk_abs i t2
+ − 973
| Abs (s, T, t') => Abs (s, T, mk_abs (i + 1) t')
+ − 974
| Bound j => if i = j then error "make_inst" else t
+ − 975
| _ => t);
+ − 976
in (f, Abs ("x", T, mk_abs 0 g)) end;
+ − 977
*}
+ − 978
+ − 979
ML {*
+ − 980
fun make_inst3 lhs t =
+ − 981
let
+ − 982
val _ $ (Abs (_, _, (_ $ ((f as Var (_, Type ("fun", [T, _]))) $ u)))) = lhs;
602
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 983
val _ $ (Abs (_, _, (_ $ g))) = t;
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 984
fun mk_abs i t =
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 985
if incr_boundvars i u aconv t then Bound i
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 986
else (case t of
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 987
t1 $ t2 => mk_abs i t1 $ mk_abs i t2
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 988
| Abs (s, T, t') => Abs (s, T, mk_abs (i + 1) t')
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 989
| Bound j => if i = j then error "make_inst" else t
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 990
| _ => t);
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 991
in (f, Abs ("x", T, mk_abs 0 g)) end;
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 992
*}
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 993
597
+ − 994
ML {*
+ − 995
fun lambda_prs_simple_conv ctxt ctrm =
+ − 996
case (term_of ctrm) of
602
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 997
((Const (@{const_name fun_map}, _) $ r1 $ a2) $ (Abs _)) =>
608
+ − 998
(let
597
+ − 999
val thy = ProofContext.theory_of ctxt
+ − 1000
val (ty_b, ty_a) = dest_fun_type (fastype_of r1)
+ − 1001
val (ty_c, ty_d) = dest_fun_type (fastype_of a2)
+ − 1002
val tyinst = map (SOME o (ctyp_of thy)) [ty_a, ty_b, ty_c, ty_d]
+ − 1003
val tinst = [NONE, NONE, SOME (cterm_of thy r1), NONE, SOME (cterm_of thy a2)]
+ − 1004
val lpi = Drule.instantiate' tyinst tinst @{thm lambda_prs}
+ − 1005
val te = @{thm eq_reflection} OF [solve_quotient_assums ctxt (solve_quotient_assums ctxt lpi)]
602
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1006
val ti =
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1007
(let
614
+ − 1008
val ts = MetaSimplifier.rewrite_rule (id_simps_get ctxt) te
602
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1009
val (insp, inst) = make_inst (term_of (Thm.lhs_of ts)) (term_of ctrm)
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1010
in
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1011
Drule.instantiate ([], [(cterm_of thy insp, cterm_of thy inst)]) ts
608
+ − 1012
end handle _ => (* TODO handle only Bind | Error "make_inst" *)
602
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1013
let
631
+ − 1014
val ts = MetaSimplifier.rewrite_rule (id_simps_get ctxt) te
+ − 1015
val _ = tracing ("ts rule:\n" ^ (Syntax.string_of_term ctxt (prop_of ts)));
+ − 1016
val _ = tracing ("redex:\n" ^ (Syntax.string_of_term ctxt (term_of ctrm)));
+ − 1017
val (insp, inst) = make_inst2 (term_of (Thm.lhs_of ts)) (term_of ctrm)
+ − 1018
in
+ − 1019
Drule.instantiate ([], [(cterm_of thy insp, cterm_of thy inst)]) ts
+ − 1020
end handle _ => (* TODO handle only Bind | Error "make_inst" *)
+ − 1021
let
+ − 1022
val (insp, inst) = make_inst3 (term_of (Thm.lhs_of te)) (term_of ctrm)
602
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1023
val td = Drule.instantiate ([], [(cterm_of thy insp, cterm_of thy inst)]) te
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1024
in
614
+ − 1025
MetaSimplifier.rewrite_rule (id_simps_get ctxt) td
602
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1026
end);
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1027
val _ = if not (Term.is_Const a2 andalso fst (dest_Const a2) = @{const_name "id"}) then
597
+ − 1028
(tracing "lambda_prs";
+ − 1029
tracing ("redex:\n" ^ (Syntax.string_of_term ctxt (term_of ctrm)));
+ − 1030
tracing ("lpi rule:\n" ^ (Syntax.string_of_term ctxt (prop_of lpi)));
+ − 1031
tracing ("te rule:\n" ^ (Syntax.string_of_term ctxt (prop_of te)));
602
e56eeb9fedb3
make_inst for lambda_prs where the second quotient is not identity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1032
tracing ("ti rule:\n" ^ (Syntax.string_of_term ctxt (prop_of ti))))
597
+ − 1033
else ()
631
+ − 1034
597
+ − 1035
in
+ − 1036
Conv.rewr_conv ti ctrm
+ − 1037
end
608
+ − 1038
handle _ => Conv.all_conv ctrm)
597
+ − 1039
| _ => Conv.all_conv ctrm
+ − 1040
*}
+ − 1041
+ − 1042
ML {*
+ − 1043
val lambda_prs_conv =
+ − 1044
More_Conv.top_conv lambda_prs_simple_conv
+ − 1045
+ − 1046
fun lambda_prs_tac ctxt = CONVERSION (lambda_prs_conv ctxt)
+ − 1047
*}
+ − 1048
659
+ − 1049
(* 1. folding of definitions and preservation lemmas; *)
+ − 1050
(* and simplification with *)
+ − 1051
thm babs_prs all_prs ex_prs
+ − 1052
(* 2. unfolding of ---> in front of everything, except *)
+ − 1053
(* bound variables *)
+ − 1054
thm fun_map_simps
615
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1055
(* 3. simplification with *)
659
+ − 1056
thm Quotient_abs_rep Quotient_rel_rep id_simps
612
+ − 1057
(* 4. Test for refl *)
597
+ − 1058
652
d8f07b5bcfae
implemented cleaning strategy with fun_map.simps on non-bounded variables; still a few rough edges
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1059
ML {*
597
+ − 1060
fun clean_tac lthy =
+ − 1061
let
+ − 1062
val thy = ProofContext.theory_of lthy;
+ − 1063
val defs = map (Thm.varifyT o symmetric o #def) (qconsts_dest thy)
615
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1064
(* FIXME: why is the Thm.varifyT needed: example where it fails is LamEx *)
658
+ − 1065
656
c86a47d4966e
Temporarily repeated fun_map_tac 4 times. Cleaning for all examples work.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1066
val thms1 = defs @ (prs_rules_get lthy) @ @{thms babs_prs all_prs ex_prs}
658
+ − 1067
val thms2 = @{thms Quotient_abs_rep Quotient_rel_rep} @ (id_simps_get lthy)
597
+ − 1068
fun simps thms = (mk_minimal_ss lthy) addsimps thms addSolver quotient_solver
+ − 1069
in
648
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1070
EVERY' [simp_tac (simps thms1),
658
+ − 1071
fun_map_tac lthy,
648
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1072
lambda_prs_tac lthy,
597
+ − 1073
simp_tac (simps thms2),
+ − 1074
TRY o rtac refl]
+ − 1075
end
+ − 1076
*}
+ − 1077
612
+ − 1078
section {* Tactic for genralisation of free variables in a goal *}
597
+ − 1079
+ − 1080
ML {*
+ − 1081
fun inst_spec ctrm =
+ − 1082
Drule.instantiate' [SOME (ctyp_of_term ctrm)] [NONE, SOME ctrm] @{thm spec}
+ − 1083
+ − 1084
fun inst_spec_tac ctrms =
+ − 1085
EVERY' (map (dtac o inst_spec) ctrms)
+ − 1086
+ − 1087
fun all_list xs trm =
+ − 1088
fold (fn (x, T) => fn t' => HOLogic.mk_all (x, T, t')) xs trm
+ − 1089
+ − 1090
fun apply_under_Trueprop f =
+ − 1091
HOLogic.dest_Trueprop #> f #> HOLogic.mk_Trueprop
+ − 1092
+ − 1093
fun gen_frees_tac ctxt =
+ − 1094
SUBGOAL (fn (concl, i) =>
+ − 1095
let
+ − 1096
val thy = ProofContext.theory_of ctxt
+ − 1097
val vrs = Term.add_frees concl []
+ − 1098
val cvrs = map (cterm_of thy o Free) vrs
+ − 1099
val concl' = apply_under_Trueprop (all_list vrs) concl
+ − 1100
val goal = Logic.mk_implies (concl', concl)
+ − 1101
val rule = Goal.prove ctxt [] [] goal
+ − 1102
(K (EVERY1 [inst_spec_tac (rev cvrs), atac]))
+ − 1103
in
+ − 1104
rtac rule i
+ − 1105
end)
+ − 1106
*}
+ − 1107
616
+ − 1108
section {* General Shape of the Lifting Procedure *}
597
+ − 1109
+ − 1110
(* - A is the original raw theorem *)
+ − 1111
(* - B is the regularized theorem *)
+ − 1112
(* - C is the rep/abs injected version of B *)
+ − 1113
(* - D is the lifted theorem *)
+ − 1114
(* *)
+ − 1115
(* - b is the regularization step *)
+ − 1116
(* - c is the rep/abs injection step *)
+ − 1117
(* - d is the cleaning part *)
+ − 1118
+ − 1119
lemma lifting_procedure:
+ − 1120
assumes a: "A"
606
+ − 1121
and b: "A \<longrightarrow> B"
597
+ − 1122
and c: "B = C"
+ − 1123
and d: "C = D"
+ − 1124
shows "D"
612
+ − 1125
using a b c d
+ − 1126
by simp
597
+ − 1127
+ − 1128
ML {*
+ − 1129
fun lift_match_error ctxt fun_str rtrm qtrm =
+ − 1130
let
+ − 1131
val rtrm_str = Syntax.string_of_term ctxt rtrm
+ − 1132
val qtrm_str = Syntax.string_of_term ctxt qtrm
+ − 1133
val msg = [enclose "[" "]" fun_str, "The quotient theorem\n", qtrm_str,
+ − 1134
"and the lifted theorem\n", rtrm_str, "do not match"]
+ − 1135
in
+ − 1136
error (space_implode " " msg)
+ − 1137
end
+ − 1138
*}
+ − 1139
+ − 1140
ML {*
+ − 1141
fun procedure_inst ctxt rtrm qtrm =
+ − 1142
let
+ − 1143
val thy = ProofContext.theory_of ctxt
+ − 1144
val rtrm' = HOLogic.dest_Trueprop rtrm
+ − 1145
val qtrm' = HOLogic.dest_Trueprop qtrm
+ − 1146
val reg_goal =
+ − 1147
Syntax.check_term ctxt (regularize_trm ctxt rtrm' qtrm')
+ − 1148
handle (LIFT_MATCH s) => lift_match_error ctxt s rtrm qtrm
+ − 1149
val inj_goal =
+ − 1150
Syntax.check_term ctxt (inj_repabs_trm ctxt (reg_goal, qtrm'))
+ − 1151
handle (LIFT_MATCH s) => lift_match_error ctxt s rtrm qtrm
+ − 1152
in
+ − 1153
Drule.instantiate' []
+ − 1154
[SOME (cterm_of thy rtrm'),
+ − 1155
SOME (cterm_of thy reg_goal),
+ − 1156
SOME (cterm_of thy inj_goal)] @{thm lifting_procedure}
+ − 1157
end
+ − 1158
*}
+ − 1159
+ − 1160
ML {*
616
+ − 1161
(* the tactic leaves three subgoals to be proved *)
597
+ − 1162
fun procedure_tac ctxt rthm =
+ − 1163
ObjectLogic.full_atomize_tac
+ − 1164
THEN' gen_frees_tac ctxt
610
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1165
THEN' CSUBGOAL (fn (goal, i) =>
597
+ − 1166
let
+ − 1167
val rthm' = atomize_thm rthm
610
2bee5ca44ef5
removed "global" data and lookup functions; had to move a tactic out from the inj_repabs_match tactic since apply_rsp interferes with a trans2 rule for ===>
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1168
val rule = procedure_inst ctxt (prop_of rthm') (term_of goal)
612
+ − 1169
val bare_goal = snd (Thm.dest_comb goal)
+ − 1170
val quot_weak = Drule.instantiate' [] [SOME bare_goal] @{thm QUOT_TRUE_i}
597
+ − 1171
in
612
+ − 1172
(rtac rule THEN' RANGE [rtac rthm', K all_tac, rtac quot_weak]) i
597
+ − 1173
end)
+ − 1174
*}
+ − 1175
616
+ − 1176
(* automatic proofs *)
612
+ − 1177
ML {*
616
+ − 1178
fun SOLVES' tac = tac THEN_ALL_NEW (K no_tac)
615
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1179
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1180
fun WARN (tac, msg) i st =
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1181
case Seq.pull ((SOLVES' tac) i st) of
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1182
NONE => (warning msg; Seq.single st)
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1183
| seqcell => Seq.make (fn () => seqcell)
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1184
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1185
fun RANGE_WARN xs = RANGE (map WARN xs)
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1186
*}
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1187
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1188
ML {*
612
+ − 1189
fun lift_tac ctxt rthm =
+ − 1190
procedure_tac ctxt rthm
615
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1191
THEN' RANGE_WARN [(regularize_tac ctxt, "Regularize proof failed."),
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1192
(all_inj_repabs_tac ctxt, "Injection proof failed."),
386a6b1a5203
the lift_tac produces a warning message if one of the three automatic proofs fails
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1193
(clean_tac ctxt, "Cleaning proof failed.")]
612
+ − 1194
*}
+ − 1195
636
+ − 1196
(* methods / interface *)
632
+ − 1197
ML {*
+ − 1198
(* FIXME: if the argument order changes then this can be just one function *)
+ − 1199
fun mk_method1 tac thm ctxt =
+ − 1200
SIMPLE_METHOD (HEADGOAL (tac ctxt thm))
+ − 1201
+ − 1202
fun mk_method2 tac ctxt =
+ − 1203
SIMPLE_METHOD (HEADGOAL (tac ctxt))
+ − 1204
*}
+ − 1205
+ − 1206
method_setup lifting =
637
+ − 1207
{* Attrib.thm >> (mk_method1 lift_tac) *}
632
+ − 1208
{* Lifting of theorems to quotient types. *}
+ − 1209
+ − 1210
method_setup lifting_setup =
637
+ − 1211
{* Attrib.thm >> (mk_method1 procedure_tac) *}
632
+ − 1212
{* Sets up the three goals for the lifting procedure. *}
+ − 1213
+ − 1214
method_setup regularize =
+ − 1215
{* Scan.succeed (mk_method2 regularize_tac) *}
+ − 1216
{* Proves automatically the regularization goals from the lifting procedure. *}
+ − 1217
+ − 1218
method_setup injection =
+ − 1219
{* Scan.succeed (mk_method2 all_inj_repabs_tac) *}
+ − 1220
{* Proves automatically the rep/abs injection goals from the lifting procedure. *}
+ − 1221
+ − 1222
method_setup cleaning =
+ − 1223
{* Scan.succeed (mk_method2 clean_tac) *}
+ − 1224
{* Proves automatically the cleaning goals from the lifting procedure. *}
+ − 1225
597
+ − 1226
end
+ − 1227