0
+ − 1
theory QuotMain
6
+ − 2
imports QuotScript QuotList Prove
71
+ − 3
uses ("quotient.ML")
0
+ − 4
begin
+ − 5
+ − 6
locale QUOT_TYPE =
+ − 7
fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+ − 8
and Abs :: "('a \<Rightarrow> bool) \<Rightarrow> 'b"
+ − 9
and Rep :: "'b \<Rightarrow> ('a \<Rightarrow> bool)"
+ − 10
assumes equiv: "EQUIV R"
+ − 11
and rep_prop: "\<And>y. \<exists>x. Rep y = R x"
+ − 12
and rep_inverse: "\<And>x. Abs (Rep x) = x"
+ − 13
and abs_inverse: "\<And>x. (Rep (Abs (R x))) = (R x)"
+ − 14
and rep_inject: "\<And>x y. (Rep x = Rep y) = (x = y)"
15
+ − 15
begin
0
+ − 16
+ − 17
definition
200
+ − 18
ABS::"'a \<Rightarrow> 'b"
+ − 19
where
0
+ − 20
"ABS x \<equiv> Abs (R x)"
+ − 21
+ − 22
definition
200
+ − 23
REP::"'b \<Rightarrow> 'a"
+ − 24
where
0
+ − 25
"REP a = Eps (Rep a)"
+ − 26
15
+ − 27
lemma lem9:
0
+ − 28
shows "R (Eps (R x)) = R x"
+ − 29
proof -
+ − 30
have a: "R x x" using equiv by (simp add: EQUIV_REFL_SYM_TRANS REFL_def)
+ − 31
then have "R x (Eps (R x))" by (rule someI)
15
+ − 32
then show "R (Eps (R x)) = R x"
0
+ − 33
using equiv unfolding EQUIV_def by simp
+ − 34
qed
+ − 35
+ − 36
theorem thm10:
24
+ − 37
shows "ABS (REP a) \<equiv> a"
+ − 38
apply (rule eq_reflection)
+ − 39
unfolding ABS_def REP_def
0
+ − 40
proof -
15
+ − 41
from rep_prop
0
+ − 42
obtain x where eq: "Rep a = R x" by auto
+ − 43
have "Abs (R (Eps (Rep a))) = Abs (R (Eps (R x)))" using eq by simp
+ − 44
also have "\<dots> = Abs (R x)" using lem9 by simp
+ − 45
also have "\<dots> = Abs (Rep a)" using eq by simp
+ − 46
also have "\<dots> = a" using rep_inverse by simp
+ − 47
finally
+ − 48
show "Abs (R (Eps (Rep a))) = a" by simp
+ − 49
qed
+ − 50
15
+ − 51
lemma REP_refl:
0
+ − 52
shows "R (REP a) (REP a)"
+ − 53
unfolding REP_def
+ − 54
by (simp add: equiv[simplified EQUIV_def])
+ − 55
+ − 56
lemma lem7:
22
+ − 57
shows "(R x = R y) = (Abs (R x) = Abs (R y))"
0
+ − 58
apply(rule iffI)
+ − 59
apply(simp)
+ − 60
apply(drule rep_inject[THEN iffD2])
+ − 61
apply(simp add: abs_inverse)
+ − 62
done
15
+ − 63
0
+ − 64
theorem thm11:
+ − 65
shows "R r r' = (ABS r = ABS r')"
+ − 66
unfolding ABS_def
+ − 67
by (simp only: equiv[simplified EQUIV_def] lem7)
+ − 68
4
+ − 69
2
+ − 70
lemma REP_ABS_rsp:
4
+ − 71
shows "R f (REP (ABS g)) = R f g"
+ − 72
and "R (REP (ABS g)) f = R g f"
23
+ − 73
by (simp_all add: thm10 thm11)
4
+ − 74
0
+ − 75
lemma QUOTIENT:
+ − 76
"QUOTIENT R ABS REP"
+ − 77
apply(unfold QUOTIENT_def)
+ − 78
apply(simp add: thm10)
+ − 79
apply(simp add: REP_refl)
+ − 80
apply(subst thm11[symmetric])
+ − 81
apply(simp add: equiv[simplified EQUIV_def])
+ − 82
done
+ − 83
21
+ − 84
lemma R_trans:
49
+ − 85
assumes ab: "R a b"
+ − 86
and bc: "R b c"
22
+ − 87
shows "R a c"
21
+ − 88
proof -
+ − 89
have tr: "TRANS R" using equiv EQUIV_REFL_SYM_TRANS[of R] by simp
+ − 90
moreover have ab: "R a b" by fact
+ − 91
moreover have bc: "R b c" by fact
22
+ − 92
ultimately show "R a c" unfolding TRANS_def by blast
21
+ − 93
qed
+ − 94
+ − 95
lemma R_sym:
49
+ − 96
assumes ab: "R a b"
22
+ − 97
shows "R b a"
21
+ − 98
proof -
+ − 99
have re: "SYM R" using equiv EQUIV_REFL_SYM_TRANS[of R] by simp
22
+ − 100
then show "R b a" using ab unfolding SYM_def by blast
21
+ − 101
qed
+ − 102
49
+ − 103
lemma R_trans2:
+ − 104
assumes ac: "R a c"
22
+ − 105
and bd: "R b d"
21
+ − 106
shows "R a b = R c d"
200
+ − 107
using ac bd
+ − 108
by (blast intro: R_trans R_sym)
21
+ − 109
+ − 110
lemma REPS_same:
25
+ − 111
shows "R (REP a) (REP b) \<equiv> (a = b)"
38
+ − 112
proof -
+ − 113
have "R (REP a) (REP b) = (a = b)"
+ − 114
proof
+ − 115
assume as: "R (REP a) (REP b)"
+ − 116
from rep_prop
+ − 117
obtain x y
+ − 118
where eqs: "Rep a = R x" "Rep b = R y" by blast
+ − 119
from eqs have "R (Eps (R x)) (Eps (R y))" using as unfolding REP_def by simp
+ − 120
then have "R x (Eps (R y))" using lem9 by simp
+ − 121
then have "R (Eps (R y)) x" using R_sym by blast
+ − 122
then have "R y x" using lem9 by simp
+ − 123
then have "R x y" using R_sym by blast
+ − 124
then have "ABS x = ABS y" using thm11 by simp
+ − 125
then have "Abs (Rep a) = Abs (Rep b)" using eqs unfolding ABS_def by simp
+ − 126
then show "a = b" using rep_inverse by simp
+ − 127
next
+ − 128
assume ab: "a = b"
+ − 129
have "REFL R" using equiv EQUIV_REFL_SYM_TRANS[of R] by simp
+ − 130
then show "R (REP a) (REP b)" unfolding REFL_def using ab by auto
+ − 131
qed
+ − 132
then show "R (REP a) (REP b) \<equiv> (a = b)" by simp
21
+ − 133
qed
+ − 134
0
+ − 135
end
+ − 136
127
b054cf6bd179
the command "quotient" can now define more than one quotient at the same time; quotients need to be separated by and
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 137
0
+ − 138
section {* type definition for the quotient type *}
+ − 139
71
+ − 140
use "quotient.ML"
0
+ − 141
185
+ − 142
declare [[map list = (map, LIST_REL)]]
+ − 143
declare [[map * = (prod_fun, prod_rel)]]
+ − 144
declare [[map "fun" = (fun_map, FUN_REL)]]
+ − 145
+ − 146
ML {* maps_lookup @{theory} "List.list" *}
+ − 147
ML {* maps_lookup @{theory} "*" *}
+ − 148
ML {* maps_lookup @{theory} "fun" *}
174
+ − 149
218
+ − 150
text {* FIXME: auxiliary function *}
193
+ − 151
ML {*
14
+ − 152
val no_vars = Thm.rule_attribute (fn context => fn th =>
+ − 153
let
+ − 154
val ctxt = Variable.set_body false (Context.proof_of context);
+ − 155
val ((_, [th']), _) = Variable.import true [th] ctxt;
+ − 156
in th' end);
+ − 157
*}
+ − 158
0
+ − 159
section {* lifting of constants *}
+ − 160
+ − 161
ML {*
218
+ − 162
fun lookup_snd _ [] _ = NONE
+ − 163
| lookup_snd eq ((value, key)::xs) key' =
+ − 164
if eq (key', key) then SOME value
+ − 165
else lookup_snd eq xs key';
+ − 166
+ − 167
fun lookup_qenv qenv qty =
254
+ − 168
(case (AList.lookup (op =) qenv qty) of
218
+ − 169
SOME rty => SOME (qty, rty)
+ − 170
| NONE => NONE)
185
+ − 171
*}
+ − 172
+ − 173
ML {*
118
+ − 174
(* calculates the aggregate abs and rep functions for a given type;
+ − 175
repF is for constants' arguments; absF is for constants;
+ − 176
function types need to be treated specially, since repF and absF
163
+ − 177
change
118
+ − 178
*)
46
+ − 179
datatype flag = absF | repF
0
+ − 180
109
386671ef36bd
fixed the problem with function types; but only type_of works; cterm_of does not work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 181
fun negF absF = repF
386671ef36bd
fixed the problem with function types; but only type_of works; cterm_of does not work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 182
| negF repF = absF
386671ef36bd
fixed the problem with function types; but only type_of works; cterm_of does not work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 183
218
+ − 184
fun get_fun flag qenv lthy ty =
0
+ − 185
let
218
+ − 186
0
+ − 187
fun get_fun_aux s fs_tys =
+ − 188
let
+ − 189
val (fs, tys) = split_list fs_tys
15
+ − 190
val (otys, ntys) = split_list tys
0
+ − 191
val oty = Type (s, otys)
+ − 192
val nty = Type (s, ntys)
+ − 193
val ftys = map (op -->) tys
+ − 194
in
130
+ − 195
(case (maps_lookup (ProofContext.theory_of lthy) s) of
110
+ − 196
SOME info => (list_comb (Const (#mapfun info, ftys ---> (oty --> nty)), fs), (oty, nty))
+ − 197
| NONE => raise ERROR ("no map association for type " ^ s))
+ − 198
end
+ − 199
118
+ − 200
fun get_fun_fun fs_tys =
110
+ − 201
let
+ − 202
val (fs, tys) = split_list fs_tys
+ − 203
val ([oty1, oty2], [nty1, nty2]) = split_list tys
128
+ − 204
val oty = nty1 --> oty2
+ − 205
val nty = oty1 --> nty2
110
+ − 206
val ftys = map (op -->) tys
+ − 207
in
118
+ − 208
(list_comb (Const (@{const_name "fun_map"}, ftys ---> oty --> nty), fs), (oty, nty))
0
+ − 209
end
+ − 210
218
+ − 211
fun get_const flag (qty, rty) =
+ − 212
let
+ − 213
val thy = ProofContext.theory_of lthy
+ − 214
val qty_name = Long_Name.base_name (fst (dest_Type qty))
+ − 215
in
+ − 216
case flag of
+ − 217
absF => (Const (Sign.full_bname thy ("ABS_" ^ qty_name), rty --> qty), (rty, qty))
+ − 218
| repF => (Const (Sign.full_bname thy ("REP_" ^ qty_name), qty --> rty), (qty, rty))
+ − 219
end
41
+ − 220
70
f3cbda066c3a
consistent usage of rty (for the raw, unquotient type); tuned a bit the Isar
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 221
fun mk_identity ty = Abs ("", ty, Bound 0)
41
+ − 222
0
+ − 223
in
254
+ − 224
if (AList.defined (op =) qenv ty)
218
+ − 225
then (get_const flag (the (lookup_qenv qenv ty)))
0
+ − 226
else (case ty of
41
+ − 227
TFree _ => (mk_identity ty, (ty, ty))
118
+ − 228
| Type (_, []) => (mk_identity ty, (ty, ty))
109
386671ef36bd
fixed the problem with function types; but only type_of works; cterm_of does not work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 229
| Type ("fun" , [ty1, ty2]) =>
218
+ − 230
get_fun_fun [get_fun (negF flag) qenv lthy ty1, get_fun flag qenv lthy ty2]
+ − 231
| Type (s, tys) => get_fun_aux s (map (get_fun flag qenv lthy) tys)
41
+ − 232
| _ => raise ERROR ("no type variables")
0
+ − 233
)
+ − 234
end
+ − 235
*}
+ − 236
180
+ − 237
41
+ − 238
text {* produces the definition for a lifted constant *}
86
+ − 239
2
+ − 240
ML {*
218
+ − 241
fun get_const_def nconst otrm qenv lthy =
0
+ − 242
let
+ − 243
val ty = fastype_of nconst
+ − 244
val (arg_tys, res_ty) = strip_type ty
14
+ − 245
218
+ − 246
val rep_fns = map (fst o get_fun repF qenv lthy) arg_tys
+ − 247
val abs_fn = (fst o get_fun absF qenv lthy) res_ty
0
+ − 248
180
+ − 249
fun mk_fun_map (t1,t2) = Const (@{const_name "fun_map"}, dummyT) $ t1 $ t2
+ − 250
218
+ − 251
val fns = Library.foldr mk_fun_map (rep_fns, abs_fn)
180
+ − 252
|> Syntax.check_term lthy
0
+ − 253
in
218
+ − 254
fns $ otrm
0
+ − 255
end
+ − 256
*}
+ − 257
218
+ − 258
0
+ − 259
ML {*
218
+ − 260
fun exchange_ty qenv ty =
254
+ − 261
case (lookup_snd (op =) qenv ty) of
218
+ − 262
SOME qty => qty
+ − 263
| NONE =>
0
+ − 264
(case ty of
218
+ − 265
Type (s, tys) => Type (s, map (exchange_ty qenv) tys)
41
+ − 266
| _ => ty
+ − 267
)
0
+ − 268
*}
+ − 269
254
+ − 270
ML {*
+ − 271
fun ppair lthy (ty1, ty2) =
+ − 272
"(" ^ (Syntax.string_of_typ lthy ty1) ^ "," ^ (Syntax.string_of_typ lthy ty2) ^ ")"
+ − 273
*}
+ − 274
+ − 275
ML {*
+ − 276
fun print_env qenv lthy =
+ − 277
map (ppair lthy) qenv
+ − 278
|> commas
+ − 279
|> tracing
+ − 280
*}
+ − 281
0
+ − 282
ML {*
218
+ − 283
fun make_const_def nconst_bname otrm mx qenv lthy =
0
+ − 284
let
218
+ − 285
val otrm_ty = fastype_of otrm
+ − 286
val nconst_ty = exchange_ty qenv otrm_ty
17
+ − 287
val nconst = Const (Binding.name_of nconst_bname, nconst_ty)
218
+ − 288
val def_trm = get_const_def nconst otrm qenv lthy
254
+ − 289
+ − 290
val _ = print_env qenv lthy
+ − 291
val _ = Syntax.string_of_typ lthy otrm_ty |> tracing
+ − 292
val _ = Syntax.string_of_typ lthy nconst_ty |> tracing
+ − 293
val _ = Syntax.string_of_term lthy def_trm |> tracing
0
+ − 294
in
79
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 295
define (nconst_bname, mx, def_trm) lthy
15
+ − 296
end
0
+ − 297
*}
+ − 298
218
+ − 299
ML {*
+ − 300
fun build_qenv lthy qtys =
+ − 301
let
+ − 302
val qenv = map (fn {qtyp, rtyp, ...} => (qtyp, rtyp)) (quotdata_lookup lthy)
254
+ − 303
val tsig = Sign.tsig_of (ProofContext.theory_of lthy)
+ − 304
+ − 305
fun find_assoc ((qty', rty')::rest) qty =
+ − 306
let
+ − 307
val inst = Type.typ_match tsig (qty', qty) Vartab.empty
+ − 308
in
+ − 309
(Envir.norm_type inst qty, Envir.norm_type inst rty')
+ − 310
end
+ − 311
| find_assoc [] qty =
+ − 312
let
+ − 313
val qtystr = quote (Syntax.string_of_typ lthy qty)
+ − 314
in
+ − 315
error (implode ["Quotient type ", qtystr, " does not exists"])
+ − 316
end
218
+ − 317
in
254
+ − 318
map (find_assoc qenv) qtys
218
+ − 319
end
+ − 320
*}
+ − 321
+ − 322
ML {*
+ − 323
val qd_parser =
+ − 324
(Args.parens (OuterParse.$$$ "for" |-- (Scan.repeat OuterParse.typ))) --
254
+ − 325
(SpecParse.constdecl -- (SpecParse.opt_thm_name ":" -- OuterParse.prop))
218
+ − 326
*}
+ − 327
+ − 328
ML {*
254
+ − 329
fun parse_qd_spec (qtystrs, ((bind, tystr, mx), (attr__, propstr))) lthy =
218
+ − 330
let
+ − 331
val qtys = map (Syntax.check_typ lthy o Syntax.parse_typ lthy) qtystrs
+ − 332
val qenv = build_qenv lthy qtys
254
+ − 333
val qty = Syntax.parse_typ lthy (the tystr) |> Syntax.check_typ lthy
218
+ − 334
val prop = Syntax.parse_prop lthy propstr |> Syntax.check_prop lthy
+ − 335
val (lhs, rhs) = Logic.dest_equals prop
+ − 336
in
+ − 337
make_const_def bind rhs mx qenv lthy |> snd
+ − 338
end
+ − 339
*}
+ − 340
+ − 341
ML {*
+ − 342
val _ = OuterSyntax.local_theory "quotient_def" "lifted definition of constants"
+ − 343
OuterKeyword.thy_decl (qd_parser >> parse_qd_spec)
+ − 344
*}
+ − 345
254
+ − 346
(* ML {* show_all_types := true *} *)
+ − 347
139
+ − 348
section {* ATOMIZE *}
+ − 349
+ − 350
text {*
+ − 351
Unabs_def converts a definition given as
+ − 352
+ − 353
c \<equiv> %x. %y. f x y
+ − 354
+ − 355
to a theorem of the form
+ − 356
+ − 357
c x y \<equiv> f x y
+ − 358
+ − 359
This function is needed to rewrite the right-hand
+ − 360
side to the left-hand side.
+ − 361
*}
+ − 362
+ − 363
ML {*
+ − 364
fun unabs_def ctxt def =
+ − 365
let
+ − 366
val (lhs, rhs) = Thm.dest_equals (cprop_of def)
+ − 367
val xs = strip_abs_vars (term_of rhs)
+ − 368
val (_, ctxt') = Variable.add_fixes (map fst xs) ctxt
+ − 369
+ − 370
val thy = ProofContext.theory_of ctxt'
+ − 371
val cxs = map (cterm_of thy o Free) xs
+ − 372
val new_lhs = Drule.list_comb (lhs, cxs)
+ − 373
+ − 374
fun get_conv [] = Conv.rewr_conv def
+ − 375
| get_conv (x::xs) = Conv.fun_conv (get_conv xs)
+ − 376
in
+ − 377
get_conv xs new_lhs |>
+ − 378
singleton (ProofContext.export ctxt' ctxt)
+ − 379
end
+ − 380
*}
+ − 381
+ − 382
lemma atomize_eqv[atomize]:
+ − 383
shows "(Trueprop A \<equiv> Trueprop B) \<equiv> (A \<equiv> B)"
+ − 384
proof
+ − 385
assume "A \<equiv> B"
+ − 386
then show "Trueprop A \<equiv> Trueprop B" by unfold
+ − 387
next
+ − 388
assume *: "Trueprop A \<equiv> Trueprop B"
+ − 389
have "A = B"
+ − 390
proof (cases A)
+ − 391
case True
+ − 392
have "A" by fact
+ − 393
then show "A = B" using * by simp
+ − 394
next
+ − 395
case False
+ − 396
have "\<not>A" by fact
+ − 397
then show "A = B" using * by auto
+ − 398
qed
+ − 399
then show "A \<equiv> B" by (rule eq_reflection)
+ − 400
qed
+ − 401
+ − 402
ML {*
+ − 403
fun atomize_thm thm =
+ − 404
let
221
+ − 405
val thm' = Thm.freezeT (forall_intr_vars thm)
139
+ − 406
val thm'' = ObjectLogic.atomize (cprop_of thm')
+ − 407
in
221
+ − 408
@{thm Pure.equal_elim_rule1} OF [thm'', thm']
139
+ − 409
end
+ − 410
*}
+ − 411
140
+ − 412
ML {* atomize_thm @{thm list.induct} *}
139
+ − 413
+ − 414
section {* REGULARIZE *}
+ − 415
+ − 416
text {* tyRel takes a type and builds a relation that a quantifier over this
+ − 417
type needs to respect. *}
+ − 418
ML {*
257
+ − 419
fun matches (ty1, ty2) =
+ − 420
Type.raw_instance (ty1, Logic.varifyT ty2);
+ − 421
139
+ − 422
fun tyRel ty rty rel lthy =
257
+ − 423
if matches (rty, ty)
139
+ − 424
then rel
+ − 425
else (case ty of
+ − 426
Type (s, tys) =>
+ − 427
let
+ − 428
val tys_rel = map (fn ty => ty --> ty --> @{typ bool}) tys;
+ − 429
val ty_out = ty --> ty --> @{typ bool};
+ − 430
val tys_out = tys_rel ---> ty_out;
+ − 431
in
+ − 432
(case (maps_lookup (ProofContext.theory_of lthy) s) of
+ − 433
SOME (info) => list_comb (Const (#relfun info, tys_out), map (fn ty => tyRel ty rty rel lthy) tys)
+ − 434
| NONE => HOLogic.eq_const ty
+ − 435
)
+ − 436
end
+ − 437
| _ => HOLogic.eq_const ty)
+ − 438
*}
+ − 439
+ − 440
definition
+ − 441
Babs :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
+ − 442
where
+ − 443
"(x \<in> p) \<Longrightarrow> (Babs p m x = m x)"
+ − 444
(* TODO: Consider defining it with an "if"; sth like:
+ − 445
Babs p m = \<lambda>x. if x \<in> p then m x else undefined
+ − 446
*)
+ − 447
+ − 448
ML {*
+ − 449
fun needs_lift (rty as Type (rty_s, _)) ty =
+ − 450
case ty of
+ − 451
Type (s, tys) =>
+ − 452
(s = rty_s) orelse (exists (needs_lift rty) tys)
+ − 453
| _ => false
+ − 454
+ − 455
*}
+ − 456
+ − 457
ML {*
+ − 458
fun mk_babs ty ty' = Const (@{const_name "Babs"}, [ty' --> @{typ bool}, ty] ---> ty)
+ − 459
fun mk_ball ty = Const (@{const_name "Ball"}, [ty, ty] ---> @{typ bool})
+ − 460
fun mk_bex ty = Const (@{const_name "Bex"}, [ty, ty] ---> @{typ bool})
+ − 461
fun mk_resp ty = Const (@{const_name Respects}, [[ty, ty] ---> @{typ bool}, ty] ---> @{typ bool})
+ − 462
*}
+ − 463
+ − 464
(* applies f to the subterm of an abstractions, otherwise to the given term *)
+ − 465
ML {*
+ − 466
fun apply_subt f trm =
+ − 467
case trm of
145
+ − 468
Abs (x, T, t) =>
+ − 469
let
+ − 470
val (x', t') = Term.dest_abs (x, T, t)
+ − 471
in
+ − 472
Term.absfree (x', T, f t')
+ − 473
end
139
+ − 474
| _ => f trm
+ − 475
*}
+ − 476
+ − 477
(* FIXME: if there are more than one quotient, then you have to look up the relation *)
+ − 478
ML {*
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 479
fun my_reg lthy rel rty trm =
139
+ − 480
case trm of
+ − 481
Abs (x, T, t) =>
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 482
if (needs_lift rty T) then
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 483
let
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 484
val rrel = tyRel T rty rel lthy
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 485
in
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 486
(mk_babs (fastype_of trm) T) $ (mk_resp T $ rrel) $ (apply_subt (my_reg lthy rel rty) trm)
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 487
end
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 488
else
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 489
Abs(x, T, (apply_subt (my_reg lthy rel rty) t))
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 490
| Const (@{const_name "All"}, ty) $ (t as Abs (x, T, _)) =>
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 491
let
139
+ − 492
val ty1 = domain_type ty
+ − 493
val ty2 = domain_type ty1
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 494
val rrel = tyRel T rty rel lthy
139
+ − 495
in
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 496
if (needs_lift rty T) then
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 497
(mk_ball ty1) $ (mk_resp ty2 $ rrel) $ (apply_subt (my_reg lthy rel rty) t)
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 498
else
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 499
Const (@{const_name "All"}, ty) $ apply_subt (my_reg lthy rel rty) t
139
+ − 500
end
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 501
| Const (@{const_name "Ex"}, ty) $ (t as Abs (x, T, _)) =>
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 502
let
139
+ − 503
val ty1 = domain_type ty
+ − 504
val ty2 = domain_type ty1
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 505
val rrel = tyRel T rty rel lthy
139
+ − 506
in
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 507
if (needs_lift rty T) then
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 508
(mk_bex ty1) $ (mk_resp ty2 $ rrel) $ (apply_subt (my_reg lthy rel rty) t)
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 509
else
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 510
Const (@{const_name "Ex"}, ty) $ apply_subt (my_reg lthy rel rty) t
139
+ − 511
end
251
+ − 512
| Const (@{const_name "op ="}, ty) $ t =>
+ − 513
if needs_lift rty (fastype_of t) then
+ − 514
(tyRel (fastype_of t) rty rel lthy) $ t
+ − 515
else Const (@{const_name "op ="}, ty) $ (my_reg lthy rel rty t)
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 516
| t1 $ t2 => (my_reg lthy rel rty t1) $ (my_reg lthy rel rty t2)
139
+ − 517
| _ => trm
+ − 518
*}
+ − 519
141
+ − 520
text {* Assumes that the given theorem is atomized *}
140
+ − 521
ML {*
+ − 522
fun build_regularize_goal thm rty rel lthy =
+ − 523
Logic.mk_implies
+ − 524
((prop_of thm),
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 525
(my_reg lthy rel rty (prop_of thm)))
140
+ − 526
*}
+ − 527
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 528
lemma universal_twice: "(\<And>x. (P x \<longrightarrow> Q x)) \<Longrightarrow> ((\<forall>x. P x) \<longrightarrow> (\<forall>x. Q x))"
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 529
apply (auto)
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 530
done
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 531
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 532
lemma implication_twice: "(c \<longrightarrow> a) \<Longrightarrow> (a \<Longrightarrow> b \<longrightarrow> d) \<Longrightarrow> (a \<longrightarrow> b) \<longrightarrow> (c \<longrightarrow> d)"
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 533
apply (auto)
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 534
done
251
+ − 535
187
+ − 536
ML {*
251
+ − 537
fun regularize thm rty rel rel_eqv rel_refl lthy =
187
+ − 538
let
+ − 539
val g = build_regularize_goal thm rty rel lthy;
+ − 540
fun tac ctxt =
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 541
(ObjectLogic.full_atomize_tac) THEN'
251
+ − 542
REPEAT_ALL_NEW (FIRST' [
+ − 543
rtac rel_refl,
+ − 544
atac,
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 545
rtac @{thm universal_twice},
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 546
(rtac @{thm impI} THEN' atac),
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 547
rtac @{thm implication_twice},
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 548
EqSubst.eqsubst_tac ctxt [0]
187
+ − 549
[(@{thm equiv_res_forall} OF [rel_eqv]),
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 550
(@{thm equiv_res_exists} OF [rel_eqv])],
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 551
(rtac @{thm impI} THEN' (asm_full_simp_tac (Simplifier.context ctxt HOL_ss)) THEN' rtac rel_refl),
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 552
(rtac @{thm RIGHT_RES_FORALL_REGULAR})
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 553
]);
187
+ − 554
val cthm = Goal.prove lthy [] [] g (fn x => tac (#context x) 1);
+ − 555
in
+ − 556
cthm OF [thm]
+ − 557
end
+ − 558
*}
+ − 559
140
+ − 560
section {* RepAbs injection *}
139
+ − 561
161
+ − 562
(* Needed to have a meta-equality *)
+ − 563
lemma id_def_sym: "(\<lambda>x. x) \<equiv> id"
+ − 564
by (simp add: id_def)
+ − 565
139
+ − 566
ML {*
218
+ − 567
fun old_exchange_ty rty qty ty =
+ − 568
if ty = rty
+ − 569
then qty
+ − 570
else
+ − 571
(case ty of
+ − 572
Type (s, tys) => Type (s, map (old_exchange_ty rty qty) tys)
+ − 573
| _ => ty
+ − 574
)
+ − 575
*}
+ − 576
+ − 577
ML {*
+ − 578
fun old_get_fun flag rty qty lthy ty =
+ − 579
get_fun flag [(qty, rty)] lthy ty
+ − 580
+ − 581
fun old_make_const_def nconst_bname otrm mx rty qty lthy =
+ − 582
make_const_def nconst_bname otrm mx [(qty, rty)] lthy
+ − 583
*}
+ − 584
235
+ − 585
text {* Does the same as 'subst' in a given prop or theorem *}
+ − 586
ML {*
+ − 587
fun eqsubst_prop ctxt thms t =
+ − 588
let
+ − 589
val goalstate = Goal.init (cterm_of (ProofContext.theory_of ctxt) t)
+ − 590
val a' = case (SINGLE (EqSubst.eqsubst_tac ctxt [0] thms 1) goalstate) of
+ − 591
NONE => error "eqsubst_prop"
+ − 592
| SOME th => cprem_of th 1
+ − 593
in term_of a' end
+ − 594
*}
+ − 595
+ − 596
ML {*
+ − 597
fun repeat_eqsubst_prop ctxt thms t =
+ − 598
repeat_eqsubst_prop ctxt thms (eqsubst_prop ctxt thms t)
+ − 599
handle _ => t
+ − 600
*}
+ − 601
+ − 602
+ − 603
ML {*
+ − 604
fun eqsubst_thm ctxt thms thm =
+ − 605
let
+ − 606
val goalstate = Goal.init (Thm.cprop_of thm)
+ − 607
val a' = case (SINGLE (EqSubst.eqsubst_tac ctxt [0] thms 1) goalstate) of
+ − 608
NONE => error "eqsubst_thm"
+ − 609
| SOME th => cprem_of th 1
+ − 610
val tac = (EqSubst.eqsubst_tac ctxt [0] thms 1) THEN simp_tac HOL_ss 1
+ − 611
val cgoal = cterm_of (ProofContext.theory_of ctxt) (Logic.mk_equals (term_of (Thm.cprop_of thm), term_of a'))
+ − 612
val rt = Toplevel.program (fn () => Goal.prove_internal [] cgoal (fn _ => tac));
+ − 613
in
+ − 614
@{thm Pure.equal_elim_rule1} OF [rt,thm]
+ − 615
end
+ − 616
*}
+ − 617
+ − 618
ML {*
+ − 619
fun repeat_eqsubst_thm ctxt thms thm =
+ − 620
repeat_eqsubst_thm ctxt thms (eqsubst_thm ctxt thms thm)
+ − 621
handle _ => thm
+ − 622
*}
+ − 623
218
+ − 624
ML {*
140
+ − 625
fun build_repabs_term lthy thm constructors rty qty =
139
+ − 626
let
+ − 627
fun mk_rep tm =
+ − 628
let
218
+ − 629
val ty = old_exchange_ty rty qty (fastype_of tm)
+ − 630
in fst (old_get_fun repF rty qty lthy ty) $ tm end
139
+ − 631
+ − 632
fun mk_abs tm =
+ − 633
let
218
+ − 634
val ty = old_exchange_ty rty qty (fastype_of tm) in
+ − 635
fst (old_get_fun absF rty qty lthy ty) $ tm end
139
+ − 636
+ − 637
fun is_constructor (Const (x, _)) = member (op =) constructors x
+ − 638
| is_constructor _ = false;
+ − 639
+ − 640
fun build_aux lthy tm =
+ − 641
case tm of
+ − 642
Abs (a as (_, vty, _)) =>
+ − 643
let
+ − 644
val (vs, t) = Term.dest_abs a;
+ − 645
val v = Free(vs, vty);
+ − 646
val t' = lambda v (build_aux lthy t)
+ − 647
in
+ − 648
if (not (needs_lift rty (fastype_of tm))) then t'
+ − 649
else mk_rep (mk_abs (
+ − 650
if not (needs_lift rty vty) then t'
+ − 651
else
+ − 652
let
+ − 653
val v' = mk_rep (mk_abs v);
+ − 654
val t1 = Envir.beta_norm (t' $ v')
+ − 655
in
+ − 656
lambda v t1
+ − 657
end
+ − 658
))
+ − 659
end
+ − 660
| x =>
+ − 661
let
+ − 662
val (opp, tms0) = Term.strip_comb tm
+ − 663
val tms = map (build_aux lthy) tms0
+ − 664
val ty = fastype_of tm
+ − 665
in
+ − 666
if (((fst (Term.dest_Const opp)) = @{const_name Respects}) handle _ => false)
+ − 667
then (list_comb (opp, (hd tms0) :: (tl tms)))
+ − 668
else if (is_constructor opp andalso needs_lift rty ty) then
+ − 669
mk_rep (mk_abs (list_comb (opp,tms)))
+ − 670
else if ((Term.is_Free opp) andalso (length tms > 0) andalso (needs_lift rty ty)) then
149
+ − 671
mk_rep(mk_abs(list_comb(opp,tms)))
139
+ − 672
else if tms = [] then opp
+ − 673
else list_comb(opp, tms)
+ − 674
end
+ − 675
in
235
+ − 676
repeat_eqsubst_prop lthy @{thms id_def_sym}
161
+ − 677
(build_aux lthy (Thm.prop_of thm))
139
+ − 678
end
+ − 679
*}
+ − 680
141
+ − 681
text {* Assumes that it is given a regularized theorem *}
139
+ − 682
ML {*
140
+ − 683
fun build_repabs_goal ctxt thm cons rty qty =
+ − 684
Logic.mk_equals ((Thm.prop_of thm), (build_repabs_term ctxt thm cons rty qty))
139
+ − 685
*}
+ − 686
187
+ − 687
ML {*
+ − 688
fun instantiate_tac thm = Subgoal.FOCUS (fn {concl, ...} =>
+ − 689
let
+ − 690
val pat = Drule.strip_imp_concl (cprop_of thm)
+ − 691
val insts = Thm.match (pat, concl)
+ − 692
in
+ − 693
rtac (Drule.instantiate insts thm) 1
152
53277fbb2dba
Simplified the proof with some tactic... Still hangs sometimes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 694
end
187
+ − 695
handle _ => no_tac
+ − 696
)
+ − 697
*}
+ − 698
+ − 699
ML {*
+ − 700
fun res_forall_rsp_tac ctxt =
+ − 701
(simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms FUN_REL.simps}))
+ − 702
THEN' rtac @{thm allI} THEN' rtac @{thm allI} THEN' rtac @{thm impI}
+ − 703
THEN' instantiate_tac @{thm RES_FORALL_RSP} ctxt THEN'
+ − 704
(simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms FUN_REL.simps}))
+ − 705
*}
+ − 706
+ − 707
ML {*
+ − 708
fun res_exists_rsp_tac ctxt =
+ − 709
(simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms FUN_REL.simps}))
+ − 710
THEN' rtac @{thm allI} THEN' rtac @{thm allI} THEN' rtac @{thm impI}
+ − 711
THEN' instantiate_tac @{thm RES_EXISTS_RSP} ctxt THEN'
+ − 712
(simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms FUN_REL.simps}))
+ − 713
*}
+ − 714
+ − 715
+ − 716
ML {*
+ − 717
fun quotient_tac quot_thm =
+ − 718
REPEAT_ALL_NEW (FIRST' [
+ − 719
rtac @{thm FUN_QUOTIENT},
+ − 720
rtac quot_thm,
+ − 721
rtac @{thm IDENTITY_QUOTIENT}
+ − 722
])
+ − 723
*}
+ − 724
+ − 725
ML {*
+ − 726
fun LAMBDA_RES_TAC ctxt i st =
+ − 727
(case (term_of o #concl o fst) (Subgoal.focus ctxt i st) of
+ − 728
(_ $ (_ $ (Abs(_,_,_))$(Abs(_,_,_)))) =>
+ − 729
(EqSubst.eqsubst_tac ctxt [0] @{thms FUN_REL.simps}) THEN'
+ − 730
(rtac @{thm allI}) THEN' (rtac @{thm allI}) THEN' (rtac @{thm impI})
+ − 731
| _ => fn _ => no_tac) i st
+ − 732
*}
+ − 733
+ − 734
ML {*
+ − 735
fun WEAK_LAMBDA_RES_TAC ctxt i st =
+ − 736
(case (term_of o #concl o fst) (Subgoal.focus ctxt i st) of
+ − 737
(_ $ (_ $ _$(Abs(_,_,_)))) =>
+ − 738
(EqSubst.eqsubst_tac ctxt [0] @{thms FUN_REL.simps}) THEN'
+ − 739
(rtac @{thm allI}) THEN' (rtac @{thm allI}) THEN' (rtac @{thm impI})
+ − 740
| (_ $ (_ $ (Abs(_,_,_))$_)) =>
+ − 741
(EqSubst.eqsubst_tac ctxt [0] @{thms FUN_REL.simps}) THEN'
+ − 742
(rtac @{thm allI}) THEN' (rtac @{thm allI}) THEN' (rtac @{thm impI})
+ − 743
| _ => fn _ => no_tac) i st
+ − 744
*}
+ − 745
206
+ − 746
ML {*
+ − 747
fun APPLY_RSP_TAC rty = Subgoal.FOCUS (fn {concl, ...} =>
+ − 748
let
+ − 749
val (_ $ (R $ (f $ _) $ (_ $ _))) = term_of concl;
+ − 750
val pat = Drule.strip_imp_concl (cprop_of @{thm APPLY_RSP});
+ − 751
val insts = Thm.match (pat, concl)
+ − 752
in
+ − 753
if needs_lift rty (type_of f) then
+ − 754
rtac (Drule.instantiate insts @{thm APPLY_RSP}) 1
+ − 755
else no_tac
+ − 756
end
+ − 757
handle _ => no_tac)
+ − 758
*}
187
+ − 759
+ − 760
ML {*
206
+ − 761
fun r_mk_comb_tac ctxt rty quot_thm reflex_thm trans_thm rsp_thms =
187
+ − 762
(FIRST' [
+ − 763
rtac @{thm FUN_QUOTIENT},
+ − 764
rtac quot_thm,
+ − 765
rtac @{thm IDENTITY_QUOTIENT},
+ − 766
rtac trans_thm,
+ − 767
LAMBDA_RES_TAC ctxt,
+ − 768
res_forall_rsp_tac ctxt,
+ − 769
res_exists_rsp_tac ctxt,
+ − 770
(
+ − 771
(simp_tac ((Simplifier.context ctxt HOL_ss) addsimps rsp_thms))
+ − 772
THEN_ALL_NEW (fn _ => no_tac)
+ − 773
),
+ − 774
(instantiate_tac @{thm REP_ABS_RSP(1)} ctxt THEN' (RANGE [quotient_tac quot_thm])),
+ − 775
rtac refl,
206
+ − 776
(* rtac @{thm arg_cong2[of _ _ _ _ "op ="]},*)
+ − 777
(APPLY_RSP_TAC rty ctxt THEN' (RANGE [quotient_tac quot_thm, quotient_tac quot_thm])),
+ − 778
Cong_Tac.cong_tac @{thm cong},
+ − 779
rtac @{thm ext},
187
+ − 780
rtac reflex_thm,
+ − 781
atac,
+ − 782
(
+ − 783
(simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms FUN_REL.simps}))
+ − 784
THEN_ALL_NEW (fn _ => no_tac)
+ − 785
),
+ − 786
WEAK_LAMBDA_RES_TAC ctxt
+ − 787
])
+ − 788
*}
+ − 789
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 790
ML {*
210
+ − 791
fun repabs lthy thm constructors rty qty quot_thm reflex_thm trans_thm rsp_thms =
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 792
let
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 793
val rt = build_repabs_term lthy thm constructors rty qty;
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 794
val rg = Logic.mk_equals ((Thm.prop_of thm), rt);
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 795
fun tac ctxt = (ObjectLogic.full_atomize_tac) THEN'
206
+ − 796
(REPEAT_ALL_NEW (r_mk_comb_tac ctxt rty quot_thm reflex_thm trans_thm rsp_thms));
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 797
val cthm = Goal.prove lthy [] [] rg (fn x => tac (#context x) 1);
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 798
in
210
+ − 799
@{thm Pure.equal_elim_rule1} OF [cthm, thm]
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 800
end
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 801
*}
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 802
187
+ − 803
section {* Cleaning the goal *}
+ − 804
236
+ − 805
lemma prod_fun_id: "prod_fun id id = id"
+ − 806
apply (simp add: prod_fun_def)
+ − 807
done
+ − 808
lemma map_id: "map id x = x"
+ − 809
apply (induct x)
+ − 810
apply (simp_all add: map.simps)
+ − 811
done
+ − 812
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 813
text {* expects atomized definition *}
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 814
ML {*
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 815
fun add_lower_defs_aux ctxt thm =
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 816
let
209
+ − 817
val e1 = @{thm fun_cong} OF [thm];
+ − 818
val f = eqsubst_thm ctxt @{thms fun_map.simps} e1;
+ − 819
val g = MetaSimplifier.rewrite_rule @{thms id_def_sym} f;
236
+ − 820
val h = repeat_eqsubst_thm ctxt @{thms FUN_MAP_I id_apply prod_fun_id map_id} g
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 821
in
236
+ − 822
thm :: (add_lower_defs_aux ctxt h)
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 823
end
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 824
handle _ => [thm]
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 825
*}
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 826
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 827
ML {*
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 828
fun add_lower_defs ctxt defs =
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 829
let
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 830
val defs_pre_sym = map symmetric defs
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 831
val defs_atom = map atomize_thm defs_pre_sym
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 832
val defs_all = flat (map (add_lower_defs_aux ctxt) defs_atom)
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 833
in
214
+ − 834
map Thm.varifyT defs_all
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 835
end
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 836
*}
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 837
191
+ − 838
ML {*
241
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 839
fun findabs_all rty tm =
191
+ − 840
case tm of
+ − 841
Abs(_, T, b) =>
+ − 842
let
+ − 843
val b' = subst_bound ((Free ("x", T)), b);
241
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 844
val tys = findabs_all rty b'
191
+ − 845
val ty = fastype_of tm
+ − 846
in if needs_lift rty ty then (ty :: tys) else tys
+ − 847
end
241
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 848
| f $ a => (findabs_all rty f) @ (findabs_all rty a)
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 849
| _ => [];
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 850
fun findabs rty tm = distinct (op =) (findabs_all rty tm)
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 851
*}
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 852
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 853
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 854
ML {*
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 855
fun findaps_all rty tm =
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 856
case tm of
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 857
Abs(_, T, b) =>
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 858
findaps_all rty (subst_bound ((Free ("x", T)), b))
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 859
| (f $ a) => (findaps_all rty f @ findaps_all rty a)
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 860
| Free (_, (T as (Type ("fun", (_ :: _))))) => (if needs_lift rty T then [T] else [])
241
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 861
| _ => [];
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 862
fun findaps rty tm = distinct (op =) (findaps_all rty tm)
191
+ − 863
*}
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 864
197
+ − 865
ML {*
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 866
fun make_simp_prs_thm lthy quot_thm thm typ =
197
+ − 867
let
+ − 868
val (_, [lty, rty]) = dest_Type typ;
+ − 869
val thy = ProofContext.theory_of lthy;
+ − 870
val (lcty, rcty) = (ctyp_of thy lty, ctyp_of thy rty)
+ − 871
val inst = [SOME lcty, NONE, SOME rcty];
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 872
val lpi = Drule.instantiate' inst [] thm;
197
+ − 873
val tac =
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 874
(compose_tac (false, lpi, 2)) THEN_ALL_NEW
197
+ − 875
(quotient_tac quot_thm);
+ − 876
val t = Goal.prove lthy [] [] (concl_of lpi) (fn _ => tac 1);
+ − 877
in
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 878
MetaSimplifier.rewrite_rule [@{thm eq_reflection} OF @{thms id_apply}] t
197
+ − 879
end
+ − 880
*}
+ − 881
+ − 882
ML {*
+ − 883
fun simp_allex_prs lthy quot thm =
+ − 884
let
+ − 885
val rwf = @{thm FORALL_PRS} OF [quot];
+ − 886
val rwfs = @{thm "HOL.sym"} OF [rwf];
+ − 887
val rwe = @{thm EXISTS_PRS} OF [quot];
+ − 888
val rwes = @{thm "HOL.sym"} OF [rwe]
+ − 889
in
+ − 890
(simp_allex_prs lthy quot (eqsubst_thm lthy [rwfs, rwes] thm))
+ − 891
end
+ − 892
handle _ => thm
+ − 893
*}
+ − 894
198
+ − 895
ML {*
239
+ − 896
fun lookup_quot_data lthy qty =
+ − 897
let
257
+ − 898
val SOME quotdata = find_first (fn x => matches ((#qtyp x), qty)) (quotdata_lookup lthy)
+ − 899
val rty = Logic.unvarifyT (#rtyp quotdata)
239
+ − 900
val rel = #rel quotdata
+ − 901
val rel_eqv = #equiv_thm quotdata
+ − 902
val rel_refl_pre = @{thm EQUIV_REFL} OF [rel_eqv]
+ − 903
val rel_refl = @{thm spec} OF [MetaSimplifier.rewrite_rule [@{thm REFL_def}] rel_refl_pre]
+ − 904
in
+ − 905
(rty, rel, rel_refl, rel_eqv)
+ − 906
end
+ − 907
*}
+ − 908
+ − 909
ML {*
+ − 910
fun lookup_quot_thms lthy qty_name =
+ − 911
let
+ − 912
val thy = ProofContext.theory_of lthy;
+ − 913
val trans2 = PureThy.get_thm thy ("QUOT_TYPE_I_" ^ qty_name ^ ".R_trans2")
+ − 914
val reps_same = PureThy.get_thm thy ("QUOT_TYPE_I_" ^ qty_name ^ ".REPS_same")
+ − 915
val quot = PureThy.get_thm thy ("QUOTIENT_" ^ qty_name)
+ − 916
in
+ − 917
(trans2, reps_same, quot)
+ − 918
end
+ − 919
*}
+ − 920
+ − 921
ML {*
+ − 922
fun lookup_quot_consts defs =
+ − 923
let
+ − 924
fun dest_term (a $ b) = (a, b);
+ − 925
val def_terms = map (snd o Logic.dest_equals o concl_of) defs;
+ − 926
in
+ − 927
map (fst o dest_Const o snd o dest_term) def_terms
+ − 928
end
+ − 929
*}
+ − 930
+ − 931
ML {*
+ − 932
fun lift_thm lthy qty qty_name rsp_thms defs t = let
+ − 933
val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data lthy qty;
+ − 934
val (trans2, reps_same, quot) = lookup_quot_thms lthy qty_name;
+ − 935
val consts = lookup_quot_consts defs;
198
+ − 936
val t_a = atomize_thm t;
251
+ − 937
val t_r = regularize t_a rty rel rel_eqv rel_refl lthy;
210
+ − 938
val t_t = repabs lthy t_r consts rty qty quot rel_refl trans2 rsp_thms;
198
+ − 939
val abs = findabs rty (prop_of t_a);
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 940
val aps = findaps rty (prop_of t_a);
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 941
val app_prs_thms = map (make_simp_prs_thm lthy quot @{thm APP_PRS}) aps;
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 942
val lam_prs_thms = map (make_simp_prs_thm lthy quot @{thm LAMBDA_PRS}) abs;
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 943
val t_l = repeat_eqsubst_thm lthy (lam_prs_thms @ app_prs_thms) t_t;
198
+ − 944
val t_a = simp_allex_prs lthy quot t_l;
239
+ − 945
val defs_sym = add_lower_defs lthy defs;
+ − 946
val t_d = repeat_eqsubst_thm lthy defs_sym t_a;
198
+ − 947
val t_r = MetaSimplifier.rewrite_rule [reps_same] t_d;
+ − 948
in
+ − 949
ObjectLogic.rulify t_r
187
+ − 950
end
198
+ − 951
*}
+ − 952
239
+ − 953
198
+ − 954
end
239
+ − 955