0
+ − 1
theory QuotMain
6
+ − 2
imports QuotScript QuotList Prove
71
+ − 3
uses ("quotient.ML")
0
+ − 4
begin
+ − 5
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
+ − 6
0
+ − 7
locale QUOT_TYPE =
+ − 8
fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+ − 9
and Abs :: "('a \<Rightarrow> bool) \<Rightarrow> 'b"
+ − 10
and Rep :: "'b \<Rightarrow> ('a \<Rightarrow> bool)"
+ − 11
assumes equiv: "EQUIV R"
+ − 12
and rep_prop: "\<And>y. \<exists>x. Rep y = R x"
+ − 13
and rep_inverse: "\<And>x. Abs (Rep x) = x"
+ − 14
and abs_inverse: "\<And>x. (Rep (Abs (R x))) = (R x)"
+ − 15
and rep_inject: "\<And>x y. (Rep x = Rep y) = (x = y)"
15
+ − 16
begin
0
+ − 17
+ − 18
definition
+ − 19
"ABS x \<equiv> Abs (R x)"
+ − 20
+ − 21
definition
+ − 22
"REP a = Eps (Rep a)"
+ − 23
15
+ − 24
lemma lem9:
0
+ − 25
shows "R (Eps (R x)) = R x"
+ − 26
proof -
+ − 27
have a: "R x x" using equiv by (simp add: EQUIV_REFL_SYM_TRANS REFL_def)
+ − 28
then have "R x (Eps (R x))" by (rule someI)
15
+ − 29
then show "R (Eps (R x)) = R x"
0
+ − 30
using equiv unfolding EQUIV_def by simp
+ − 31
qed
+ − 32
+ − 33
theorem thm10:
24
+ − 34
shows "ABS (REP a) \<equiv> a"
+ − 35
apply (rule eq_reflection)
+ − 36
unfolding ABS_def REP_def
0
+ − 37
proof -
15
+ − 38
from rep_prop
0
+ − 39
obtain x where eq: "Rep a = R x" by auto
+ − 40
have "Abs (R (Eps (Rep a))) = Abs (R (Eps (R x)))" using eq by simp
+ − 41
also have "\<dots> = Abs (R x)" using lem9 by simp
+ − 42
also have "\<dots> = Abs (Rep a)" using eq by simp
+ − 43
also have "\<dots> = a" using rep_inverse by simp
+ − 44
finally
+ − 45
show "Abs (R (Eps (Rep a))) = a" by simp
+ − 46
qed
+ − 47
15
+ − 48
lemma REP_refl:
0
+ − 49
shows "R (REP a) (REP a)"
+ − 50
unfolding REP_def
+ − 51
by (simp add: equiv[simplified EQUIV_def])
+ − 52
+ − 53
lemma lem7:
22
+ − 54
shows "(R x = R y) = (Abs (R x) = Abs (R y))"
0
+ − 55
apply(rule iffI)
+ − 56
apply(simp)
+ − 57
apply(drule rep_inject[THEN iffD2])
+ − 58
apply(simp add: abs_inverse)
+ − 59
done
15
+ − 60
0
+ − 61
theorem thm11:
+ − 62
shows "R r r' = (ABS r = ABS r')"
+ − 63
unfolding ABS_def
+ − 64
by (simp only: equiv[simplified EQUIV_def] lem7)
+ − 65
4
+ − 66
2
+ − 67
lemma REP_ABS_rsp:
4
+ − 68
shows "R f (REP (ABS g)) = R f g"
+ − 69
and "R (REP (ABS g)) f = R g f"
23
+ − 70
by (simp_all add: thm10 thm11)
4
+ − 71
0
+ − 72
lemma QUOTIENT:
+ − 73
"QUOTIENT R ABS REP"
+ − 74
apply(unfold QUOTIENT_def)
+ − 75
apply(simp add: thm10)
+ − 76
apply(simp add: REP_refl)
+ − 77
apply(subst thm11[symmetric])
+ − 78
apply(simp add: equiv[simplified EQUIV_def])
+ − 79
done
+ − 80
21
+ − 81
lemma R_trans:
49
+ − 82
assumes ab: "R a b"
+ − 83
and bc: "R b c"
22
+ − 84
shows "R a c"
21
+ − 85
proof -
+ − 86
have tr: "TRANS R" using equiv EQUIV_REFL_SYM_TRANS[of R] by simp
+ − 87
moreover have ab: "R a b" by fact
+ − 88
moreover have bc: "R b c" by fact
22
+ − 89
ultimately show "R a c" unfolding TRANS_def by blast
21
+ − 90
qed
+ − 91
+ − 92
lemma R_sym:
49
+ − 93
assumes ab: "R a b"
22
+ − 94
shows "R b a"
21
+ − 95
proof -
+ − 96
have re: "SYM R" using equiv EQUIV_REFL_SYM_TRANS[of R] by simp
22
+ − 97
then show "R b a" using ab unfolding SYM_def by blast
21
+ − 98
qed
+ − 99
49
+ − 100
lemma R_trans2:
+ − 101
assumes ac: "R a c"
22
+ − 102
and bd: "R b d"
21
+ − 103
shows "R a b = R c d"
+ − 104
proof
23
+ − 105
assume "R a b"
21
+ − 106
then have "R b a" using R_sym by blast
+ − 107
then have "R b c" using ac R_trans by blast
+ − 108
then have "R c b" using R_sym by blast
+ − 109
then show "R c d" using bd R_trans by blast
+ − 110
next
23
+ − 111
assume "R c d"
21
+ − 112
then have "R a d" using ac R_trans by blast
+ − 113
then have "R d a" using R_sym by blast
+ − 114
then have "R b a" using bd R_trans by blast
+ − 115
then show "R a b" using R_sym by blast
+ − 116
qed
+ − 117
+ − 118
lemma REPS_same:
25
+ − 119
shows "R (REP a) (REP b) \<equiv> (a = b)"
38
+ − 120
proof -
+ − 121
have "R (REP a) (REP b) = (a = b)"
+ − 122
proof
+ − 123
assume as: "R (REP a) (REP b)"
+ − 124
from rep_prop
+ − 125
obtain x y
+ − 126
where eqs: "Rep a = R x" "Rep b = R y" by blast
+ − 127
from eqs have "R (Eps (R x)) (Eps (R y))" using as unfolding REP_def by simp
+ − 128
then have "R x (Eps (R y))" using lem9 by simp
+ − 129
then have "R (Eps (R y)) x" using R_sym by blast
+ − 130
then have "R y x" using lem9 by simp
+ − 131
then have "R x y" using R_sym by blast
+ − 132
then have "ABS x = ABS y" using thm11 by simp
+ − 133
then have "Abs (Rep a) = Abs (Rep b)" using eqs unfolding ABS_def by simp
+ − 134
then show "a = b" using rep_inverse by simp
+ − 135
next
+ − 136
assume ab: "a = b"
+ − 137
have "REFL R" using equiv EQUIV_REFL_SYM_TRANS[of R] by simp
+ − 138
then show "R (REP a) (REP b)" unfolding REFL_def using ab by auto
+ − 139
qed
+ − 140
then show "R (REP a) (REP b) \<equiv> (a = b)" by simp
21
+ − 141
qed
+ − 142
0
+ − 143
end
+ − 144
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
+ − 145
0
+ − 146
section {* type definition for the quotient type *}
+ − 147
71
+ − 148
use "quotient.ML"
0
+ − 149
130
+ − 150
(* mapfuns for some standard types *)
+ − 151
setup {*
+ − 152
maps_update @{type_name "list"} {mapfun = @{const_name "map"}, relfun = @{const_name "LIST_REL"}} #>
+ − 153
maps_update @{type_name "*"} {mapfun = @{const_name "prod_fun"}, relfun = @{const_name "prod_rel"}} #>
+ − 154
maps_update @{type_name "fun"} {mapfun = @{const_name "fun_map"}, relfun = @{const_name "FUN_REL"}}
+ − 155
*}
+ − 156
+ − 157
+ − 158
ML {* maps_lookup @{theory} @{type_name list} *}
79
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 159
14
+ − 160
ML {*
+ − 161
val no_vars = Thm.rule_attribute (fn context => fn th =>
+ − 162
let
+ − 163
val ctxt = Variable.set_body false (Context.proof_of context);
+ − 164
val ((_, [th']), _) = Variable.import true [th] ctxt;
+ − 165
in th' end);
+ − 166
*}
+ − 167
0
+ − 168
section {* various tests for quotient types*}
+ − 169
datatype trm =
15
+ − 170
var "nat"
0
+ − 171
| app "trm" "trm"
+ − 172
| lam "nat" "trm"
13
+ − 173
49
+ − 174
axiomatization
+ − 175
RR :: "trm \<Rightarrow> trm \<Rightarrow> bool"
26
+ − 176
where
14
+ − 177
r_eq: "EQUIV RR"
0
+ − 178
81
+ − 179
quotient qtrm = trm / "RR"
79
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 180
apply(rule r_eq)
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 181
done
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 182
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 183
typ qtrm
0
+ − 184
term Rep_qtrm
2
+ − 185
term REP_qtrm
0
+ − 186
term Abs_qtrm
+ − 187
term ABS_qtrm
+ − 188
thm QUOT_TYPE_qtrm
+ − 189
thm QUOTIENT_qtrm
79
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 190
thm REP_qtrm_def
2
+ − 191
16
+ − 192
(* Test interpretation *)
+ − 193
thm QUOT_TYPE_I_qtrm.thm11
26
+ − 194
thm QUOT_TYPE.thm11
16
+ − 195
21
+ − 196
print_theorems
+ − 197
0
+ − 198
thm Rep_qtrm
+ − 199
+ − 200
text {* another test *}
+ − 201
datatype 'a trm' =
15
+ − 202
var' "'a"
0
+ − 203
| app' "'a trm'" "'a trm'"
+ − 204
| lam' "'a" "'a trm'"
15
+ − 205
0
+ − 206
consts R' :: "'a trm' \<Rightarrow> 'a trm' \<Rightarrow> bool"
+ − 207
axioms r_eq': "EQUIV R'"
+ − 208
81
+ − 209
quotient qtrm' = "'a trm'" / "R'"
79
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 210
apply(rule r_eq')
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 211
done
0
+ − 212
14
+ − 213
print_theorems
+ − 214
0
+ − 215
term ABS_qtrm'
+ − 216
term REP_qtrm'
+ − 217
thm QUOT_TYPE_qtrm'
+ − 218
thm QUOTIENT_qtrm'
+ − 219
thm Rep_qtrm'
+ − 220
14
+ − 221
0
+ − 222
text {* a test with lists of terms *}
+ − 223
datatype t =
+ − 224
vr "string"
+ − 225
| ap "t list"
+ − 226
| lm "string" "t"
+ − 227
+ − 228
consts Rt :: "t \<Rightarrow> t \<Rightarrow> bool"
+ − 229
axioms t_eq: "EQUIV Rt"
+ − 230
81
+ − 231
quotient qt = "t" / "Rt"
79
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 232
by (rule t_eq)
0
+ − 233
+ − 234
section {* lifting of constants *}
+ − 235
+ − 236
ML {*
118
+ − 237
(* calculates the aggregate abs and rep functions for a given type;
+ − 238
repF is for constants' arguments; absF is for constants;
+ − 239
function types need to be treated specially, since repF and absF
+ − 240
change
+ − 241
*)
46
+ − 242
datatype flag = absF | repF
0
+ − 243
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
+ − 244
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
+ − 245
| 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
+ − 246
46
+ − 247
fun get_fun flag rty qty lthy ty =
0
+ − 248
let
+ − 249
val qty_name = Long_Name.base_name (fst (dest_Type qty))
15
+ − 250
0
+ − 251
fun get_fun_aux s fs_tys =
+ − 252
let
+ − 253
val (fs, tys) = split_list fs_tys
15
+ − 254
val (otys, ntys) = split_list tys
0
+ − 255
val oty = Type (s, otys)
+ − 256
val nty = Type (s, ntys)
+ − 257
val ftys = map (op -->) tys
+ − 258
in
130
+ − 259
(case (maps_lookup (ProofContext.theory_of lthy) s) of
110
+ − 260
SOME info => (list_comb (Const (#mapfun info, ftys ---> (oty --> nty)), fs), (oty, nty))
+ − 261
| NONE => raise ERROR ("no map association for type " ^ s))
+ − 262
end
+ − 263
118
+ − 264
fun get_fun_fun fs_tys =
110
+ − 265
let
+ − 266
val (fs, tys) = split_list fs_tys
+ − 267
val ([oty1, oty2], [nty1, nty2]) = split_list tys
128
+ − 268
val oty = nty1 --> oty2
+ − 269
val nty = oty1 --> nty2
110
+ − 270
val ftys = map (op -->) tys
+ − 271
in
118
+ − 272
(list_comb (Const (@{const_name "fun_map"}, ftys ---> oty --> nty), fs), (oty, nty))
0
+ − 273
end
+ − 274
46
+ − 275
fun get_const absF = (Const ("QuotMain.ABS_" ^ qty_name, rty --> qty), (rty, qty))
+ − 276
| get_const repF = (Const ("QuotMain.REP_" ^ qty_name, qty --> rty), (qty, rty))
41
+ − 277
70
f3cbda066c3a
consistent usage of rty (for the raw, unquotient type); tuned a bit the Isar
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 278
fun mk_identity ty = Abs ("", ty, Bound 0)
41
+ − 279
0
+ − 280
in
+ − 281
if ty = qty
46
+ − 282
then (get_const flag)
0
+ − 283
else (case ty of
41
+ − 284
TFree _ => (mk_identity ty, (ty, ty))
118
+ − 285
| 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
+ − 286
| Type ("fun" , [ty1, ty2]) =>
118
+ − 287
get_fun_fun [get_fun (negF flag) rty qty lthy ty1, get_fun flag rty qty lthy ty2]
46
+ − 288
| Type (s, tys) => get_fun_aux s (map (get_fun flag rty qty lthy) tys)
41
+ − 289
| _ => raise ERROR ("no type variables")
0
+ − 290
)
+ − 291
end
+ − 292
*}
+ − 293
+ − 294
ML {*
118
+ − 295
get_fun repF @{typ t} @{typ qt} @{context} @{typ "((((qt \<Rightarrow> qt) \<Rightarrow> qt) \<Rightarrow> qt) list) * nat"}
2
+ − 296
|> fst
+ − 297
|> Syntax.string_of_term @{context}
+ − 298
|> writeln
+ − 299
*}
+ − 300
46
+ − 301
ML {*
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
+ − 302
get_fun absF @{typ t} @{typ qt} @{context} @{typ "qt * nat"}
46
+ − 303
|> fst
118
+ − 304
|> Syntax.string_of_term @{context}
+ − 305
|> writeln
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
+ − 306
*}
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
+ − 307
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
+ − 308
ML {*
118
+ − 309
get_fun absF @{typ t} @{typ qt} @{context} @{typ "(qt \<Rightarrow> qt) \<Rightarrow> qt"}
+ − 310
|> fst
+ − 311
|> Syntax.pretty_term @{context}
+ − 312
|> Pretty.string_of
+ − 313
|> writeln
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
+ − 314
*}
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
+ − 315
41
+ − 316
text {* produces the definition for a lifted constant *}
86
+ − 317
2
+ − 318
ML {*
0
+ − 319
fun get_const_def nconst oconst rty qty lthy =
+ − 320
let
+ − 321
val ty = fastype_of nconst
+ − 322
val (arg_tys, res_ty) = strip_type ty
14
+ − 323
0
+ − 324
val fresh_args = arg_tys |> map (pair "x")
15
+ − 325
|> Variable.variant_frees lthy [nconst, oconst]
0
+ − 326
|> map Free
+ − 327
46
+ − 328
val rep_fns = map (fst o get_fun repF rty qty lthy) arg_tys
+ − 329
val abs_fn = (fst o get_fun absF rty qty lthy) res_ty
0
+ − 330
+ − 331
in
+ − 332
map (op $) (rep_fns ~~ fresh_args)
+ − 333
|> curry list_comb oconst
+ − 334
|> curry (op $) abs_fn
+ − 335
|> fold_rev lambda fresh_args
+ − 336
end
+ − 337
*}
+ − 338
+ − 339
ML {*
15
+ − 340
fun exchange_ty rty qty ty =
49
+ − 341
if ty = rty
41
+ − 342
then qty
15
+ − 343
else
0
+ − 344
(case ty of
+ − 345
Type (s, tys) => Type (s, map (exchange_ty rty qty) tys)
41
+ − 346
| _ => ty
+ − 347
)
0
+ − 348
*}
+ − 349
+ − 350
ML {*
17
+ − 351
fun make_const_def nconst_bname oconst mx rty qty lthy =
0
+ − 352
let
+ − 353
val oconst_ty = fastype_of oconst
+ − 354
val nconst_ty = exchange_ty rty qty oconst_ty
17
+ − 355
val nconst = Const (Binding.name_of nconst_bname, nconst_ty)
0
+ − 356
val def_trm = get_const_def nconst oconst rty qty lthy
+ − 357
in
79
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 358
define (nconst_bname, mx, def_trm) lthy
15
+ − 359
end
0
+ − 360
*}
+ − 361
111
+ − 362
(* A test whether get_fun works properly
+ − 363
consts bla :: "(t \<Rightarrow> t) \<Rightarrow> t"
+ − 364
local_setup {*
+ − 365
fn lthy => (Toplevel.program (fn () =>
+ − 366
make_const_def @{binding bla'} @{term "bla"} NoSyn @{typ "t"} @{typ "qt"} lthy
+ − 367
)) |> snd
+ − 368
*}
+ − 369
*)
+ − 370
2
+ − 371
local_setup {*
41
+ − 372
make_const_def @{binding VR} @{term "vr"} NoSyn @{typ "t"} @{typ "qt"} #> snd #>
+ − 373
make_const_def @{binding AP} @{term "ap"} NoSyn @{typ "t"} @{typ "qt"} #> snd #>
17
+ − 374
make_const_def @{binding LM} @{term "lm"} NoSyn @{typ "t"} @{typ "qt"} #> snd
2
+ − 375
*}
+ − 376
41
+ − 377
term vr
+ − 378
term ap
+ − 379
term lm
+ − 380
thm VR_def AP_def LM_def
14
+ − 381
term LM
2
+ − 382
term VR
+ − 383
term AP
+ − 384
0
+ − 385
text {* a test with functions *}
+ − 386
datatype 'a t' =
+ − 387
vr' "string"
+ − 388
| ap' "('a t') * ('a t')"
+ − 389
| lm' "'a" "string \<Rightarrow> ('a t')"
+ − 390
+ − 391
consts Rt' :: "('a t') \<Rightarrow> ('a t') \<Rightarrow> bool"
+ − 392
axioms t_eq': "EQUIV Rt'"
+ − 393
81
+ − 394
quotient qt' = "'a t'" / "Rt'"
79
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 395
apply(rule t_eq')
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 396
done
0
+ − 397
14
+ − 398
print_theorems
2
+ − 399
0
+ − 400
local_setup {*
41
+ − 401
make_const_def @{binding VR'} @{term "vr'"} NoSyn @{typ "'a t'"} @{typ "'a qt'"} #> snd #>
+ − 402
make_const_def @{binding AP'} @{term "ap'"} NoSyn @{typ "'a t'"} @{typ "'a qt'"} #> snd #>
17
+ − 403
make_const_def @{binding LM'} @{term "lm'"} NoSyn @{typ "'a t'"} @{typ "'a qt'"} #> snd
0
+ − 404
*}
+ − 405
41
+ − 406
term vr'
+ − 407
term ap'
+ − 408
term ap'
+ − 409
thm VR'_def AP'_def LM'_def
14
+ − 410
term LM'
0
+ − 411
term VR'
+ − 412
term AP'
+ − 413
14
+ − 414
0
+ − 415
text {* finite set example *}
13
+ − 416
print_syntax
14
+ − 417
inductive
13
+ − 418
list_eq (infix "\<approx>" 50)
0
+ − 419
where
+ − 420
"a#b#xs \<approx> b#a#xs"
+ − 421
| "[] \<approx> []"
+ − 422
| "xs \<approx> ys \<Longrightarrow> ys \<approx> xs"
+ − 423
| "a#a#xs \<approx> a#xs"
+ − 424
| "xs \<approx> ys \<Longrightarrow> a#xs \<approx> a#ys"
+ − 425
| "\<lbrakk>xs1 \<approx> xs2; xs2 \<approx> xs3\<rbrakk> \<Longrightarrow> xs1 \<approx> xs3"
+ − 426
47
+ − 427
lemma list_eq_refl:
0
+ − 428
shows "xs \<approx> xs"
13
+ − 429
apply (induct xs)
+ − 430
apply (auto intro: list_eq.intros)
+ − 431
done
0
+ − 432
+ − 433
lemma equiv_list_eq:
+ − 434
shows "EQUIV list_eq"
13
+ − 435
unfolding EQUIV_REFL_SYM_TRANS REFL_def SYM_def TRANS_def
47
+ − 436
apply(auto intro: list_eq.intros list_eq_refl)
13
+ − 437
done
0
+ − 438
81
+ − 439
quotient fset = "'a list" / "list_eq"
79
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 440
apply(rule equiv_list_eq)
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 441
done
0
+ − 442
14
+ − 443
print_theorems
+ − 444
0
+ − 445
typ "'a fset"
+ − 446
thm "Rep_fset"
+ − 447
+ − 448
local_setup {*
17
+ − 449
make_const_def @{binding EMPTY} @{term "[]"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
0
+ − 450
*}
+ − 451
14
+ − 452
term Nil
0
+ − 453
term EMPTY
2
+ − 454
thm EMPTY_def
+ − 455
0
+ − 456
+ − 457
local_setup {*
17
+ − 458
make_const_def @{binding INSERT} @{term "op #"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
0
+ − 459
*}
+ − 460
14
+ − 461
term Cons
0
+ − 462
term INSERT
2
+ − 463
thm INSERT_def
0
+ − 464
+ − 465
local_setup {*
17
+ − 466
make_const_def @{binding UNION} @{term "op @"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
0
+ − 467
*}
+ − 468
15
+ − 469
term append
0
+ − 470
term UNION
2
+ − 471
thm UNION_def
+ − 472
7
+ − 473
0
+ − 474
thm QUOTIENT_fset
+ − 475
14
+ − 476
thm QUOT_TYPE_I_fset.thm11
9
+ − 477
+ − 478
13
+ − 479
fun
+ − 480
membship :: "'a \<Rightarrow> 'a list \<Rightarrow> bool" (infix "memb" 100)
0
+ − 481
where
+ − 482
m1: "(x memb []) = False"
+ − 483
| m2: "(x memb (y#xs)) = ((x=y) \<or> (x memb xs))"
+ − 484
2
+ − 485
lemma mem_respects:
9
+ − 486
fixes z
2
+ − 487
assumes a: "list_eq x y"
9
+ − 488
shows "(z memb x) = (z memb y)"
13
+ − 489
using a by induct auto
+ − 490
2
+ − 491
29
+ − 492
fun
+ − 493
card1 :: "'a list \<Rightarrow> nat"
+ − 494
where
+ − 495
card1_nil: "(card1 []) = 0"
+ − 496
| card1_cons: "(card1 (x # xs)) = (if (x memb xs) then (card1 xs) else (Suc (card1 xs)))"
+ − 497
+ − 498
local_setup {*
+ − 499
make_const_def @{binding card} @{term "card1"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
+ − 500
*}
+ − 501
+ − 502
term card1
+ − 503
term card
+ − 504
thm card_def
+ − 505
+ − 506
(* text {*
+ − 507
Maybe make_const_def should require a theorem that says that the particular lifted function
+ − 508
respects the relation. With it such a definition would be impossible:
+ − 509
make_const_def @{binding CARD} @{term "length"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
41
+ − 510
*}*)
29
+ − 511
+ − 512
lemma card1_rsp:
+ − 513
fixes a b :: "'a list"
+ − 514
assumes e: "a \<approx> b"
+ − 515
shows "card1 a = card1 b"
+ − 516
using e apply induct
+ − 517
apply (simp_all add:mem_respects)
+ − 518
done
+ − 519
+ − 520
lemma card1_0:
+ − 521
fixes a :: "'a list"
+ − 522
shows "(card1 a = 0) = (a = [])"
+ − 523
apply (induct a)
+ − 524
apply (simp)
+ − 525
apply (simp_all)
+ − 526
apply meson
+ − 527
apply (simp_all)
+ − 528
done
+ − 529
+ − 530
lemma not_mem_card1:
+ − 531
fixes x :: "'a"
+ − 532
fixes xs :: "'a list"
+ − 533
shows "~(x memb xs) \<Longrightarrow> card1 (x # xs) = Suc (card1 xs)"
+ − 534
by simp
+ − 535
+ − 536
+ − 537
lemma mem_cons:
+ − 538
fixes x :: "'a"
+ − 539
fixes xs :: "'a list"
+ − 540
assumes a : "x memb xs"
+ − 541
shows "x # xs \<approx> xs"
49
+ − 542
using a
38
+ − 543
apply (induct xs)
+ − 544
apply (auto intro: list_eq.intros)
+ − 545
done
29
+ − 546
+ − 547
lemma card1_suc:
+ − 548
fixes xs :: "'a list"
+ − 549
fixes n :: "nat"
+ − 550
assumes c: "card1 xs = Suc n"
+ − 551
shows "\<exists>a ys. ~(a memb ys) \<and> xs \<approx> (a # ys)"
49
+ − 552
using c
38
+ − 553
apply(induct xs)
+ − 554
apply (metis Suc_neq_Zero card1_0)
47
+ − 555
apply (metis QUOT_TYPE_I_fset.R_trans QuotMain.card1_cons list_eq_refl mem_cons)
38
+ − 556
done
29
+ − 557
97
+ − 558
primrec
+ − 559
fold1
+ − 560
where
+ − 561
"fold1 f (g :: 'a \<Rightarrow> 'b) (z :: 'b) [] = z"
+ − 562
| "fold1 f g z (a # A) =
+ − 563
(if ((!u v. (f u v = f v u))
+ − 564
\<and> (!u v w. ((f u (f v w) = f (f u v) w))))
+ − 565
then (
+ − 566
if (a memb A) then (fold1 f g z A) else (f (g a) (fold1 f g z A))
+ − 567
) else z)"
+ − 568
100
+ − 569
(* fold1_def is not usable, but: *)
97
+ − 570
thm fold1.simps
+ − 571
12
+ − 572
lemma cons_preserves:
+ − 573
fixes z
+ − 574
assumes a: "xs \<approx> ys"
+ − 575
shows "(z # xs) \<approx> (z # ys)"
13
+ − 576
using a by (rule QuotMain.list_eq.intros(5))
12
+ − 577
60
4826ad24f772
First theorem with quantifiers. Learned how to use sledgehammer.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 578
lemma fs1_strong_cases:
4826ad24f772
First theorem with quantifiers. Learned how to use sledgehammer.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 579
fixes X :: "'a list"
4826ad24f772
First theorem with quantifiers. Learned how to use sledgehammer.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 580
shows "(X = []) \<or> (\<exists>a. \<exists> Y. (~(a memb Y) \<and> (X \<approx> a # Y)))"
4826ad24f772
First theorem with quantifiers. Learned how to use sledgehammer.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 581
apply (induct X)
4826ad24f772
First theorem with quantifiers. Learned how to use sledgehammer.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 582
apply (simp)
72
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 583
apply (metis QUOT_TYPE_I_fset.thm11 list_eq_refl mem_cons QuotMain.m1)
60
4826ad24f772
First theorem with quantifiers. Learned how to use sledgehammer.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 584
done
29
+ − 585
10
+ − 586
local_setup {*
17
+ − 587
make_const_def @{binding IN} @{term "membship"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
10
+ − 588
*}
+ − 589
0
+ − 590
term membship
+ − 591
term IN
2
+ − 592
thm IN_def
66
+ − 593
2
+ − 594
lemmas a = QUOT_TYPE.ABS_def[OF QUOT_TYPE_fset]
+ − 595
thm QUOT_TYPE.thm11[OF QUOT_TYPE_fset, THEN iffD1, simplified a]
0
+ − 596
2
+ − 597
lemma yy:
+ − 598
shows "(False = x memb []) = (False = IN (x::nat) EMPTY)"
+ − 599
unfolding IN_def EMPTY_def
+ − 600
apply(rule_tac f="(op =) False" in arg_cong)
+ − 601
apply(rule mem_respects)
14
+ − 602
apply(simp only: QUOT_TYPE_I_fset.REP_ABS_rsp)
2
+ − 603
apply(rule list_eq.intros)
0
+ − 604
done
+ − 605
2
+ − 606
lemma
+ − 607
shows "IN (x::nat) EMPTY = False"
+ − 608
using m1
+ − 609
apply -
+ − 610
apply(rule yy[THEN iffD1, symmetric])
+ − 611
apply(simp)
0
+ − 612
done
+ − 613
2
+ − 614
lemma
+ − 615
shows "((x=y) \<or> (IN x xs) = (IN (x::nat) (INSERT y xs))) =
23
+ − 616
((x=y) \<or> x memb REP_fset xs = x memb (y # REP_fset xs))"
2
+ − 617
unfolding IN_def INSERT_def
+ − 618
apply(rule_tac f="(op \<or>) (x=y)" in arg_cong)
+ − 619
apply(rule_tac f="(op =) (x memb REP_fset xs)" in arg_cong)
+ − 620
apply(rule mem_respects)
+ − 621
apply(rule list_eq.intros(3))
+ − 622
apply(unfold REP_fset_def ABS_fset_def)
5
+ − 623
apply(simp only: QUOT_TYPE.REP_ABS_rsp[OF QUOT_TYPE_fset])
47
+ − 624
apply(rule list_eq_refl)
2
+ − 625
done
0
+ − 626
7
+ − 627
lemma append_respects_fst:
14
+ − 628
assumes a : "list_eq l1 l2"
4
+ − 629
shows "list_eq (l1 @ s) (l2 @ s)"
3
+ − 630
using a
+ − 631
apply(induct)
4
+ − 632
apply(auto intro: list_eq.intros)
47
+ − 633
apply(simp add: list_eq_refl)
3
+ − 634
done
+ − 635
18
+ − 636
lemma yyy:
3
+ − 637
shows "
+ − 638
(
+ − 639
(UNION EMPTY s = s) &
+ − 640
((UNION (INSERT e s1) s2) = (INSERT e (UNION s1 s2)))
+ − 641
) = (
+ − 642
((ABS_fset ([] @ REP_fset s)) = s) &
+ − 643
((ABS_fset ((e # (REP_fset s1)) @ REP_fset s2)) = ABS_fset (e # (REP_fset s1 @ REP_fset s2)))
+ − 644
)"
+ − 645
unfolding UNION_def EMPTY_def INSERT_def
+ − 646
apply(rule_tac f="(op &)" in arg_cong2)
+ − 647
apply(rule_tac f="(op =)" in arg_cong2)
14
+ − 648
apply(simp only: QUOT_TYPE_I_fset.thm11[symmetric])
7
+ − 649
apply(rule append_respects_fst)
18
+ − 650
apply(simp only: QUOT_TYPE_I_fset.REP_ABS_rsp)
47
+ − 651
apply(rule list_eq_refl)
3
+ − 652
apply(simp)
+ − 653
apply(rule_tac f="(op =)" in arg_cong2)
14
+ − 654
apply(simp only: QUOT_TYPE_I_fset.thm11[symmetric])
7
+ − 655
apply(rule append_respects_fst)
14
+ − 656
apply(simp only: QUOT_TYPE_I_fset.REP_ABS_rsp)
47
+ − 657
apply(rule list_eq_refl)
14
+ − 658
apply(simp only: QUOT_TYPE_I_fset.thm11[symmetric])
3
+ − 659
apply(rule list_eq.intros(5))
14
+ − 660
apply(simp only: QUOT_TYPE_I_fset.REP_ABS_rsp)
47
+ − 661
apply(rule list_eq_refl)
3
+ − 662
done
+ − 663
+ − 664
lemma
+ − 665
shows "
+ − 666
(UNION EMPTY s = s) &
7
+ − 667
((UNION (INSERT e s1) s2) = (INSERT e (UNION s1 s2)))"
+ − 668
apply (simp add: yyy)
24
+ − 669
apply (simp add: QUOT_TYPE_I_fset.thm10)
7
+ − 670
done
3
+ − 671
8
+ − 672
ML {*
55
+ − 673
fun mk_rep x = @{term REP_fset} $ x;
+ − 674
fun mk_abs x = @{term ABS_fset} $ x;
107
+ − 675
val consts = [@{const_name "Nil"}, @{const_name "Cons"},
+ − 676
@{const_name "membship"}, @{const_name "card1"},
+ − 677
@{const_name "append"}, @{const_name "fold1"}];
8
+ − 678
*}
+ − 679
107
+ − 680
ML {* val tm = @{term "x :: 'a list"} *}
+ − 681
ML {* val ty = exchange_ty @{typ "'a list"} @{typ "'a fset"} (fastype_of tm) *}
+ − 682
ML {* (fst (get_fun repF @{typ "'a list"} @{typ "'a fset"} @{context} ty)) $ tm *}
+ − 683
56
+ − 684
ML {* val qty = @{typ "'a fset"} *}
+ − 685
ML {* val tt = Type ("fun", [Type ("fun", [qty, @{typ "prop"}]), @{typ "prop"}]) *}
58
f2565006dc5a
Just one atomize is enough for the currently lifted theorems. Properly lift 'all' and 'Ex'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 686
ML {* val fall = Const(@{const_name all}, dummyT) *}
f2565006dc5a
Just one atomize is enough for the currently lifted theorems. Properly lift 'all' and 'Ex'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 687
8
+ − 688
ML {*
107
+ − 689
fun needs_lift (rty as Type (rty_s, _)) ty =
+ − 690
case ty of
+ − 691
Type (s, tys) =>
+ − 692
(s = rty_s) orelse (exists (needs_lift rty) tys)
+ − 693
| _ => false
+ − 694
+ − 695
*}
+ − 696
+ − 697
ML {*
+ − 698
fun build_goal_term lthy thm constructors rty qty =
+ − 699
let
+ − 700
fun mk_rep tm =
+ − 701
let
+ − 702
val ty = exchange_ty rty qty (fastype_of tm)
+ − 703
in fst (get_fun repF rty qty lthy ty) $ tm end
+ − 704
+ − 705
fun mk_abs tm =
+ − 706
let
+ − 707
val _ = tracing (Syntax.string_of_term @{context} tm)
+ − 708
val _ = tracing (Syntax.string_of_typ @{context} (fastype_of tm))
+ − 709
val ty = exchange_ty rty qty (fastype_of tm) in
+ − 710
fst (get_fun absF rty qty lthy ty) $ tm end
+ − 711
+ − 712
fun is_constructor (Const (x, _)) = member (op =) constructors x
+ − 713
| is_constructor _ = false;
+ − 714
+ − 715
fun build_aux lthy tm =
+ − 716
case tm of
+ − 717
Abs (a as (_, vty, _)) =>
+ − 718
let
+ − 719
val (vs, t) = Term.dest_abs a;
+ − 720
val v = Free(vs, vty);
+ − 721
val t' = lambda v (build_aux lthy t)
+ − 722
in
+ − 723
if (not (needs_lift rty (fastype_of tm))) then t'
+ − 724
else mk_rep (mk_abs (
+ − 725
if not (needs_lift rty vty) then t'
+ − 726
else
+ − 727
let
+ − 728
val v' = mk_rep (mk_abs v);
+ − 729
val t1 = Envir.beta_norm (t' $ v')
+ − 730
in
+ − 731
lambda v t1
+ − 732
end
+ − 733
))
+ − 734
end
+ − 735
| x =>
+ − 736
let
+ − 737
val _ = tracing (">>" ^ Syntax.string_of_term @{context} tm)
+ − 738
val (opp, tms0) = Term.strip_comb tm
+ − 739
val tms = map (build_aux lthy) tms0
+ − 740
val ty = fastype_of tm
+ − 741
in
+ − 742
if (((fst (Term.dest_Const opp)) = @{const_name Respects}) handle _ => false)
+ − 743
then (list_comb (opp, (hd tms0) :: (tl tms)))
+ − 744
else if (is_constructor opp andalso needs_lift rty ty) then
+ − 745
mk_rep (mk_abs (list_comb (opp,tms)))
+ − 746
else if ((Term.is_Free opp) andalso (length tms > 0) andalso (needs_lift rty ty)) then
+ − 747
mk_rep(mk_abs(list_comb(opp,tms)))
+ − 748
else if tms = [] then opp
+ − 749
else list_comb(opp, tms)
+ − 750
end
+ − 751
in
+ − 752
build_aux lthy (Thm.prop_of thm)
+ − 753
end
+ − 754
+ − 755
*}
+ − 756
+ − 757
ML {*
+ − 758
fun build_goal_term_old ctxt thm constructors rty qty mk_rep mk_abs =
49
+ − 759
let
55
+ − 760
fun mk_rep_abs x = mk_rep (mk_abs x);
+ − 761
49
+ − 762
fun is_constructor (Const (x, _)) = member (op =) constructors x
+ − 763
| is_constructor _ = false;
+ − 764
43
+ − 765
fun maybe_mk_rep_abs t =
49
+ − 766
let
+ − 767
val _ = writeln ("Maybe: " ^ Syntax.string_of_term ctxt t)
+ − 768
in
55
+ − 769
if fastype_of t = rty then mk_rep_abs t else t
49
+ − 770
end;
+ − 771
56
+ − 772
fun is_all (Const ("all", Type("fun", [Type("fun", [ty, _]), _]))) = ty = rty
+ − 773
| is_all _ = false;
+ − 774
58
f2565006dc5a
Just one atomize is enough for the currently lifted theorems. Properly lift 'all' and 'Ex'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 775
fun is_ex (Const ("Ex", Type("fun", [Type("fun", [ty, _]), _]))) = ty = rty
f2565006dc5a
Just one atomize is enough for the currently lifted theorems. Properly lift 'all' and 'Ex'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 776
| is_ex _ = false;
f2565006dc5a
Just one atomize is enough for the currently lifted theorems. Properly lift 'all' and 'Ex'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 777
49
+ − 778
fun build_aux ctxt1 tm =
+ − 779
let
+ − 780
val (head, args) = Term.strip_comb tm;
+ − 781
val args' = map (build_aux ctxt1) args;
+ − 782
in
+ − 783
(case head of
+ − 784
Abs (x, T, t) =>
55
+ − 785
if T = rty then let
+ − 786
val ([x'], ctxt2) = Variable.variant_fixes [x] ctxt1;
+ − 787
val v = Free (x', qty);
+ − 788
val t' = subst_bound (mk_rep v, t);
+ − 789
val rec_term = build_aux ctxt2 t';
+ − 790
val _ = tracing (Syntax.string_of_term ctxt2 t')
+ − 791
val _ = tracing (Syntax.string_of_term ctxt2 (Term.lambda_name (x, v) rec_term))
+ − 792
in
+ − 793
Term.lambda_name (x, v) rec_term
+ − 794
end else let
49
+ − 795
val ([x'], ctxt2) = Variable.variant_fixes [x] ctxt1;
+ − 796
val v = Free (x', T);
+ − 797
val t' = subst_bound (v, t);
+ − 798
val rec_term = build_aux ctxt2 t';
+ − 799
in Term.lambda_name (x, v) rec_term end
58
f2565006dc5a
Just one atomize is enough for the currently lifted theorems. Properly lift 'all' and 'Ex'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 800
| _ => (* I assume that we never lift 'prop' since it is not of sort type *)
f2565006dc5a
Just one atomize is enough for the currently lifted theorems. Properly lift 'all' and 'Ex'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 801
if is_all head then
f2565006dc5a
Just one atomize is enough for the currently lifted theorems. Properly lift 'all' and 'Ex'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 802
list_comb (Const(@{const_name all}, dummyT), args') |> Syntax.check_term ctxt1
f2565006dc5a
Just one atomize is enough for the currently lifted theorems. Properly lift 'all' and 'Ex'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 803
else if is_ex head then
f2565006dc5a
Just one atomize is enough for the currently lifted theorems. Properly lift 'all' and 'Ex'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 804
list_comb (Const(@{const_name Ex}, dummyT), args') |> Syntax.check_term ctxt1
56
+ − 805
else if is_constructor head then
49
+ − 806
maybe_mk_rep_abs (list_comb (head, map maybe_mk_rep_abs args'))
56
+ − 807
else
+ − 808
maybe_mk_rep_abs (list_comb (head, args'))
+ − 809
)
49
+ − 810
end;
+ − 811
in
58
f2565006dc5a
Just one atomize is enough for the currently lifted theorems. Properly lift 'all' and 'Ex'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 812
build_aux ctxt (Thm.prop_of thm)
49
+ − 813
end
+ − 814
*}
+ − 815
58
f2565006dc5a
Just one atomize is enough for the currently lifted theorems. Properly lift 'all' and 'Ex'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 816
ML {*
107
+ − 817
fun build_goal ctxt thm cons rty qty =
114
+ − 818
Logic.mk_equals ((Thm.prop_of thm), (build_goal_term ctxt thm cons rty qty))
58
f2565006dc5a
Just one atomize is enough for the currently lifted theorems. Properly lift 'all' and 'Ex'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 819
*}
f2565006dc5a
Just one atomize is enough for the currently lifted theorems. Properly lift 'all' and 'Ex'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 820
67
+ − 821
ML {*
107
+ − 822
val m1_novars = snd(no_vars ((Context.Theory @{theory}), @{thm m2}))
67
+ − 823
*}
107
+ − 824
67
+ − 825
ML {*
107
+ − 826
cterm_of @{theory} (prop_of m1_novars);
+ − 827
cterm_of @{theory} (build_goal_term @{context} m1_novars consts @{typ "'a list"} @{typ "'a fset"});
+ − 828
cterm_of @{theory} (build_goal_term_old @{context} m1_novars consts @{typ "'a list"} @{typ "'a fset"} mk_rep mk_abs)
67
+ − 829
*}
+ − 830
58
f2565006dc5a
Just one atomize is enough for the currently lifted theorems. Properly lift 'all' and 'Ex'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 831
100
+ − 832
text {*
+ − 833
Unabs_def converts a definition given as
+ − 834
+ − 835
c \<equiv> %x. %y. f x y
+ − 836
+ − 837
to a theorem of the form
+ − 838
+ − 839
c x y \<equiv> f x y
+ − 840
+ − 841
This function is needed to rewrite the right-hand
+ − 842
side to the left-hand side.
+ − 843
*}
+ − 844
+ − 845
ML {*
+ − 846
fun unabs_def ctxt def =
+ − 847
let
+ − 848
val (lhs, rhs) = Thm.dest_equals (cprop_of def)
+ − 849
val xs = strip_abs_vars (term_of rhs)
+ − 850
val (_, ctxt') = Variable.add_fixes (map fst xs) ctxt
+ − 851
+ − 852
val thy = ProofContext.theory_of ctxt'
+ − 853
val cxs = map (cterm_of thy o Free) xs
+ − 854
val new_lhs = Drule.list_comb (lhs, cxs)
+ − 855
+ − 856
fun get_conv [] = Conv.rewr_conv def
+ − 857
| get_conv (x::xs) = Conv.fun_conv (get_conv xs)
+ − 858
in
+ − 859
get_conv xs new_lhs |>
+ − 860
singleton (ProofContext.export ctxt' ctxt)
+ − 861
end
+ − 862
*}
+ − 863
+ − 864
ML {* (* Test: *) @{thm IN_def}; unabs_def @{context} @{thm IN_def} *}
+ − 865
29
+ − 866
ML {* val fset_defs = @{thms EMPTY_def IN_def UNION_def card_def INSERT_def} *}
68
e8660818c755
renamed unlam_def to unabs_def (matching the function abs_def in drule.ML)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 867
ML {* val fset_defs_sym = map (fn t => symmetric (unabs_def @{context} t)) fset_defs *}
18
+ − 868
70
f3cbda066c3a
consistent usage of rty (for the raw, unquotient type); tuned a bit the Isar
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 869
lemma atomize_eqv[atomize]:
f3cbda066c3a
consistent usage of rty (for the raw, unquotient type); tuned a bit the Isar
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 870
shows "(Trueprop A \<equiv> Trueprop B) \<equiv> (A \<equiv> B)"
49
+ − 871
proof
70
f3cbda066c3a
consistent usage of rty (for the raw, unquotient type); tuned a bit the Isar
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 872
assume "A \<equiv> B"
f3cbda066c3a
consistent usage of rty (for the raw, unquotient type); tuned a bit the Isar
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 873
then show "Trueprop A \<equiv> Trueprop B" by unfold
49
+ − 874
next
70
f3cbda066c3a
consistent usage of rty (for the raw, unquotient type); tuned a bit the Isar
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 875
assume *: "Trueprop A \<equiv> Trueprop B"
49
+ − 876
have "A = B"
+ − 877
proof (cases A)
+ − 878
case True
70
f3cbda066c3a
consistent usage of rty (for the raw, unquotient type); tuned a bit the Isar
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 879
have "A" by fact
f3cbda066c3a
consistent usage of rty (for the raw, unquotient type); tuned a bit the Isar
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 880
then show "A = B" using * by simp
49
+ − 881
next
+ − 882
case False
70
f3cbda066c3a
consistent usage of rty (for the raw, unquotient type); tuned a bit the Isar
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 883
have "\<not>A" by fact
f3cbda066c3a
consistent usage of rty (for the raw, unquotient type); tuned a bit the Isar
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 884
then show "A = B" using * by auto
49
+ − 885
qed
70
f3cbda066c3a
consistent usage of rty (for the raw, unquotient type); tuned a bit the Isar
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 886
then show "A \<equiv> B" by (rule eq_reflection)
49
+ − 887
qed
+ − 888
103
+ − 889
(* Has all the theorems about fset plugged in. These should be parameters to the tactic *)
+ − 890
ML {*
+ − 891
fun transconv_fset_tac' ctxt =
+ − 892
(LocalDefs.unfold_tac @{context} fset_defs) THEN
+ − 893
ObjectLogic.full_atomize_tac 1 THEN
+ − 894
REPEAT_ALL_NEW (FIRST' [
+ − 895
rtac @{thm list_eq_refl},
+ − 896
rtac @{thm cons_preserves},
+ − 897
rtac @{thm mem_respects},
+ − 898
rtac @{thm card1_rsp},
+ − 899
rtac @{thm QUOT_TYPE_I_fset.R_trans2},
+ − 900
CHANGED o (simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms QUOT_TYPE_I_fset.REP_ABS_rsp})),
+ − 901
Cong_Tac.cong_tac @{thm cong},
+ − 902
rtac @{thm ext}
+ − 903
]) 1
+ − 904
*}
+ − 905
+ − 906
ML {*
+ − 907
val m1_novars = snd(no_vars ((Context.Theory @{theory}), @{thm m1}))
107
+ − 908
val goal = build_goal @{context} m1_novars consts @{typ "'a list"} @{typ "'a fset"}
103
+ − 909
val cgoal = cterm_of @{theory} goal
+ − 910
val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
+ − 911
*}
+ − 912
+ − 913
(*notation ( output) "prop" ("#_" [1000] 1000) *)
+ − 914
notation ( output) "Trueprop" ("#_" [1000] 1000)
+ − 915
+ − 916
30
e35198635d64
Using "atomize" the versions with arbitrary Trueprops can be proven.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 917
prove {* (Thm.term_of cgoal2) *}
103
+ − 918
apply (tactic {* transconv_fset_tac' @{context} *})
30
e35198635d64
Using "atomize" the versions with arbitrary Trueprops can be proven.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 919
done
18
+ − 920
20
+ − 921
thm length_append (* Not true but worth checking that the goal is correct *)
+ − 922
ML {*
49
+ − 923
val m1_novars = snd(no_vars ((Context.Theory @{theory}), @{thm length_append}))
107
+ − 924
val goal = build_goal @{context} m1_novars consts @{typ "'a list"} @{typ "'a fset"}
20
+ − 925
val cgoal = cterm_of @{theory} goal
21
+ − 926
val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
20
+ − 927
*}
30
e35198635d64
Using "atomize" the versions with arbitrary Trueprops can be proven.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 928
prove {* Thm.term_of cgoal2 *}
103
+ − 929
apply (tactic {* transconv_fset_tac' @{context} *})
25
+ − 930
sorry
19
+ − 931
20
+ − 932
thm m2
+ − 933
ML {*
49
+ − 934
val m1_novars = snd(no_vars ((Context.Theory @{theory}), @{thm m2}))
107
+ − 935
val goal = build_goal @{context} m1_novars consts @{typ "'a list"} @{typ "'a fset"}
20
+ − 936
val cgoal = cterm_of @{theory} goal
21
+ − 937
val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
20
+ − 938
*}
30
e35198635d64
Using "atomize" the versions with arbitrary Trueprops can be proven.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 939
prove {* Thm.term_of cgoal2 *}
103
+ − 940
apply (tactic {* transconv_fset_tac' @{context} *})
30
e35198635d64
Using "atomize" the versions with arbitrary Trueprops can be proven.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 941
done
21
+ − 942
+ − 943
thm list_eq.intros(4)
+ − 944
ML {*
49
+ − 945
val m1_novars = snd(no_vars ((Context.Theory @{theory}), @{thm list_eq.intros(4)}))
107
+ − 946
val goal = build_goal @{context} m1_novars consts @{typ "'a list"} @{typ "'a fset"}
21
+ − 947
val cgoal = cterm_of @{theory} goal
+ − 948
val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
24
+ − 949
val cgoal3 = Thm.rhs_of (MetaSimplifier.rewrite true @{thms QUOT_TYPE_I_fset.thm10} cgoal2)
21
+ − 950
*}
+ − 951
24
+ − 952
(* It is the same, but we need a name for it. *)
52
6584b1ceedce
Partially fix lifting of card_suc. The quantified variable still needs to be changed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 953
prove zzz : {* Thm.term_of cgoal3 *}
103
+ − 954
apply (tactic {* transconv_fset_tac' @{context} *})
21
+ − 955
done
+ − 956
107
+ − 957
(*lemma zzz' :
24
+ − 958
"(REP_fset (INSERT a (INSERT a (ABS_fset xs))) \<approx> REP_fset (INSERT a (ABS_fset xs)))"
+ − 959
using list_eq.intros(4) by (simp only: zzz)
+ − 960
+ − 961
thm QUOT_TYPE_I_fset.REPS_same
25
+ − 962
ML {* val zzz'' = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.REPS_same} @{thm zzz'} *}
107
+ − 963
*)
25
+ − 964
21
+ − 965
thm list_eq.intros(5)
+ − 966
ML {*
49
+ − 967
val m1_novars = snd(no_vars ((Context.Theory @{theory}), @{thm list_eq.intros(5)}))
107
+ − 968
val goal = build_goal @{context} m1_novars consts @{typ "'a list"} @{typ "'a fset"}
103
+ − 969
val cgoal = cterm_of @{theory} goal
49
+ − 970
val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite true fset_defs_sym cgoal)
21
+ − 971
*}
30
e35198635d64
Using "atomize" the versions with arbitrary Trueprops can be proven.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 972
prove {* Thm.term_of cgoal2 *}
103
+ − 973
apply (tactic {* transconv_fset_tac' @{context} *})
21
+ − 974
done
18
+ − 975
99
+ − 976
section {* handling quantifiers - regularize *}
+ − 977
+ − 978
text {* tyRel takes a type and builds a relation that a quantifier over this
+ − 979
type needs to respect. *}
+ − 980
ML {*
+ − 981
fun tyRel ty rty rel lthy =
102
+ − 982
if ty = rty
+ − 983
then rel
99
+ − 984
else (case ty of
+ − 985
Type (s, tys) =>
+ − 986
let
+ − 987
val tys_rel = map (fn ty => ty --> ty --> @{typ bool}) tys;
+ − 988
val ty_out = ty --> ty --> @{typ bool};
+ − 989
val tys_out = tys_rel ---> ty_out;
+ − 990
in
130
+ − 991
(case (maps_lookup (ProofContext.theory_of lthy) s) of
99
+ − 992
SOME (info) => list_comb (Const (#relfun info, tys_out), map (fn ty => tyRel ty rty rel lthy) tys)
102
+ − 993
| NONE => HOLogic.eq_const ty
99
+ − 994
)
+ − 995
end
102
+ − 996
| _ => HOLogic.eq_const ty)
99
+ − 997
*}
+ − 998
+ − 999
ML {*
+ − 1000
cterm_of @{theory} (tyRel @{typ "trm \<Rightarrow> bool"} @{typ "trm"} @{term "RR"} @{context})
+ − 1001
*}
+ − 1002
+ − 1003
definition
+ − 1004
Babs :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
+ − 1005
where
+ − 1006
"(x \<in> p) \<Longrightarrow> (Babs p m x = m x)"
+ − 1007
(* TODO: Consider defining it with an "if"; sth like:
+ − 1008
Babs p m = \<lambda>x. if x \<in> p then m x else undefined
+ − 1009
*)
+ − 1010
+ − 1011
ML {*
+ − 1012
(* trm \<Rightarrow> new_trm *)
+ − 1013
fun regularise trm rty rel lthy =
+ − 1014
case trm of
+ − 1015
Abs (x, T, t) =>
+ − 1016
if (needs_lift rty T) then let
+ − 1017
val ([x'], lthy2) = Variable.variant_fixes [x] lthy;
+ − 1018
val v = Free (x', T);
+ − 1019
val t' = subst_bound (v, t);
+ − 1020
val rec_term = regularise t' rty rel lthy2;
+ − 1021
val lam_term = Term.lambda_name (x, v) rec_term;
+ − 1022
val sub_res_term = tyRel T rty rel lthy;
+ − 1023
val respects = Const (@{const_name Respects}, (fastype_of sub_res_term) --> T --> @{typ bool});
+ − 1024
val res_term = respects $ sub_res_term;
+ − 1025
val ty = fastype_of trm;
+ − 1026
val rabs = Const (@{const_name Babs}, (fastype_of res_term) --> ty --> ty);
+ − 1027
val rabs_term = (rabs $ res_term) $ lam_term;
+ − 1028
in
+ − 1029
rabs_term
+ − 1030
end else let
+ − 1031
val ([x'], lthy2) = Variable.variant_fixes [x] lthy;
+ − 1032
val v = Free (x', T);
+ − 1033
val t' = subst_bound (v, t);
+ − 1034
val rec_term = regularise t' rty rel lthy2;
+ − 1035
in
+ − 1036
Term.lambda_name (x, v) rec_term
+ − 1037
end
+ − 1038
| ((Const (@{const_name "All"}, at)) $ (Abs (x, T, t))) =>
+ − 1039
if (needs_lift rty T) then let
+ − 1040
val ([x'], lthy2) = Variable.variant_fixes [x] lthy;
+ − 1041
val v = Free (x', T);
+ − 1042
val t' = subst_bound (v, t);
+ − 1043
val rec_term = regularise t' rty rel lthy2;
+ − 1044
val lam_term = Term.lambda_name (x, v) rec_term;
+ − 1045
val sub_res_term = tyRel T rty rel lthy;
+ − 1046
val respects = Const (@{const_name Respects}, (fastype_of sub_res_term) --> T --> @{typ bool});
+ − 1047
val res_term = respects $ sub_res_term;
+ − 1048
val ty = fastype_of lam_term;
+ − 1049
val rall = Const (@{const_name Ball}, (fastype_of res_term) --> ty --> @{typ bool});
+ − 1050
val rall_term = (rall $ res_term) $ lam_term;
+ − 1051
in
+ − 1052
rall_term
+ − 1053
end else let
+ − 1054
val ([x'], lthy2) = Variable.variant_fixes [x] lthy;
+ − 1055
val v = Free (x', T);
+ − 1056
val t' = subst_bound (v, t);
+ − 1057
val rec_term = regularise t' rty rel lthy2;
+ − 1058
val lam_term = Term.lambda_name (x, v) rec_term
+ − 1059
in
+ − 1060
Const(@{const_name "All"}, at) $ lam_term
+ − 1061
end
+ − 1062
| ((Const (@{const_name "All"}, at)) $ P) =>
+ − 1063
let
+ − 1064
val (_, [al, _]) = dest_Type (fastype_of P);
+ − 1065
val ([x], lthy2) = Variable.variant_fixes [""] lthy;
+ − 1066
val v = (Free (x, al));
+ − 1067
val abs = Term.lambda_name (x, v) (P $ v);
+ − 1068
in regularise ((Const (@{const_name "All"}, at)) $ abs) rty rel lthy2 end
+ − 1069
| ((Const (@{const_name "Ex"}, at)) $ (Abs (x, T, t))) =>
+ − 1070
if (needs_lift rty T) then let
+ − 1071
val ([x'], lthy2) = Variable.variant_fixes [x] lthy;
+ − 1072
val v = Free (x', T);
+ − 1073
val t' = subst_bound (v, t);
+ − 1074
val rec_term = regularise t' rty rel lthy2;
+ − 1075
val lam_term = Term.lambda_name (x, v) rec_term;
+ − 1076
val sub_res_term = tyRel T rty rel lthy;
+ − 1077
val respects = Const (@{const_name Respects}, (fastype_of sub_res_term) --> T --> @{typ bool});
+ − 1078
val res_term = respects $ sub_res_term;
+ − 1079
val ty = fastype_of lam_term;
+ − 1080
val rall = Const (@{const_name Bex}, (fastype_of res_term) --> ty --> @{typ bool});
+ − 1081
val rall_term = (rall $ res_term) $ lam_term;
+ − 1082
in
+ − 1083
rall_term
+ − 1084
end else let
+ − 1085
val ([x'], lthy2) = Variable.variant_fixes [x] lthy;
+ − 1086
val v = Free (x', T);
+ − 1087
val t' = subst_bound (v, t);
+ − 1088
val rec_term = regularise t' rty rel lthy2;
+ − 1089
val lam_term = Term.lambda_name (x, v) rec_term
+ − 1090
in
+ − 1091
Const(@{const_name "Ex"}, at) $ lam_term
+ − 1092
end
+ − 1093
| ((Const (@{const_name "Ex"}, at)) $ P) =>
+ − 1094
let
+ − 1095
val (_, [al, _]) = dest_Type (fastype_of P);
+ − 1096
val ([x], lthy2) = Variable.variant_fixes [""] lthy;
+ − 1097
val v = (Free (x, al));
+ − 1098
val abs = Term.lambda_name (x, v) (P $ v);
+ − 1099
in regularise ((Const (@{const_name "Ex"}, at)) $ abs) rty rel lthy2 end
+ − 1100
| a $ b => (regularise a rty rel lthy) $ (regularise b rty rel lthy)
+ − 1101
| _ => trm
+ − 1102
+ − 1103
*}
+ − 1104
+ − 1105
ML {*
+ − 1106
cterm_of @{theory} (regularise @{term "\<lambda>x :: int. x"} @{typ "trm"} @{term "RR"} @{context});
+ − 1107
cterm_of @{theory} (regularise @{term "\<lambda>x :: trm. x"} @{typ "trm"} @{term "RR"} @{context});
+ − 1108
cterm_of @{theory} (regularise @{term "\<forall>x :: trm. P x"} @{typ "trm"} @{term "RR"} @{context});
+ − 1109
cterm_of @{theory} (regularise @{term "\<exists>x :: trm. P x"} @{typ "trm"} @{term "RR"} @{context});
+ − 1110
cterm_of @{theory} (regularise @{term "All (P :: trm \<Rightarrow> bool)"} @{typ "trm"} @{term "RR"} @{context});
+ − 1111
*}
+ − 1112
131
+ − 1113
(* my version of regularise *)
+ − 1114
(****************************)
+ − 1115
+ − 1116
(* some helper functions *)
134
+ − 1117
+ − 1118
131
+ − 1119
ML {*
134
+ − 1120
fun mk_babs ty ty' = Const (@{const_name "Babs"}, [ty' --> @{typ bool}, ty] ---> ty)
131
+ − 1121
fun mk_ball ty = Const (@{const_name "Ball"}, [ty, ty] ---> @{typ bool})
+ − 1122
fun mk_bex ty = Const (@{const_name "Bex"}, [ty, ty] ---> @{typ bool})
134
+ − 1123
fun mk_resp ty = Const (@{const_name Respects}, [[ty, ty] ---> @{typ bool}, ty] ---> @{typ bool})
131
+ − 1124
*}
+ − 1125
+ − 1126
(* applies f to the subterm of an abstractions, otherwise to the given term *)
+ − 1127
ML {*
+ − 1128
fun apply_subt f trm =
+ − 1129
case trm of
+ − 1130
Abs (x, T, t) => Abs (x, T, f t)
+ − 1131
| _ => f trm
+ − 1132
*}
+ − 1133
133
+ − 1134
131
+ − 1135
(* FIXME: assumes always the typ is qty! *)
133
+ − 1136
(* FIXME: if there are more than one quotient, then you have to look up the relation *)
131
+ − 1137
ML {*
+ − 1138
fun my_reg rel trm =
+ − 1139
case trm of
+ − 1140
Abs (x, T, t) =>
+ − 1141
let
+ − 1142
val ty1 = fastype_of trm
+ − 1143
in
133
+ − 1144
(mk_babs ty1 T) $ (mk_resp T $ rel) $ (Abs (x, T, my_reg rel t))
131
+ − 1145
end
133
+ − 1146
| Const (@{const_name "All"}, ty) $ t =>
131
+ − 1147
let
133
+ − 1148
val ty1 = domain_type ty
131
+ − 1149
val ty2 = domain_type ty1
+ − 1150
in
+ − 1151
(mk_ball ty1) $ (mk_resp ty2 $ rel) $ (apply_subt (my_reg rel) t)
+ − 1152
end
133
+ − 1153
| Const (@{const_name "Ex"}, ty) $ t =>
131
+ − 1154
let
133
+ − 1155
val ty1 = domain_type ty
131
+ − 1156
val ty2 = domain_type ty1
+ − 1157
in
+ − 1158
(mk_bex ty1) $ (mk_resp ty2 $ rel) $ (apply_subt (my_reg rel) t)
+ − 1159
end
132
+ − 1160
| t1 $ t2 => (my_reg rel t1) $ (my_reg rel t2)
131
+ − 1161
| _ => trm
+ − 1162
*}
+ − 1163
+ − 1164
ML {*
+ − 1165
cterm_of @{theory} (regularise @{term "\<lambda>x::trm. x"} @{typ "trm"} @{term "RR"} @{context});
+ − 1166
cterm_of @{theory} (my_reg @{term "RR"} @{term "\<lambda>x::trm. x"})
+ − 1167
*}
+ − 1168
+ − 1169
ML {*
+ − 1170
cterm_of @{theory} (regularise @{term "\<forall>(x::trm) (y::trm). P x y"} @{typ "trm"} @{term "RR"} @{context});
+ − 1171
cterm_of @{theory} (my_reg @{term "RR"} @{term "\<forall>(x::trm) (y::trm). P x y"})
+ − 1172
*}
+ − 1173
+ − 1174
ML {*
+ − 1175
cterm_of @{theory} (regularise @{term "\<forall>x::trm. P x"} @{typ "trm"} @{term "RR"} @{context});
+ − 1176
cterm_of @{theory} (my_reg @{term "RR"} @{term "\<forall>x::trm. P x"})
+ − 1177
*}
+ − 1178
+ − 1179
ML {*
+ − 1180
cterm_of @{theory} (regularise @{term "\<exists>x::trm. P x"} @{typ "trm"} @{term "RR"} @{context});
+ − 1181
cterm_of @{theory} (my_reg @{term "RR"} @{term "\<exists>x::trm. P x"})
+ − 1182
*}
+ − 1183
134
+ − 1184
(* my version is not eta-expanded, but that should be OK *)
131
+ − 1185
ML {*
+ − 1186
cterm_of @{theory} (regularise @{term "All (P::trm \<Rightarrow> bool)"} @{typ "trm"} @{term "RR"} @{context});
+ − 1187
cterm_of @{theory} (my_reg @{term "RR"} @{term "All (P::trm \<Rightarrow> bool)"})
+ − 1188
*}
+ − 1189
99
+ − 1190
ML {*
+ − 1191
fun atomize_thm thm =
+ − 1192
let
+ − 1193
val thm' = forall_intr_vars thm
+ − 1194
val thm'' = ObjectLogic.atomize (cprop_of thm')
+ − 1195
in
+ − 1196
Simplifier.rewrite_rule [thm''] thm'
+ − 1197
end
+ − 1198
*}
+ − 1199
+ − 1200
thm list.induct
114
+ − 1201
lemma list_induct_hol4:
+ − 1202
fixes P :: "'a list \<Rightarrow> bool"
+ − 1203
assumes "((P []) \<and> (\<forall>t. (P t) \<longrightarrow> (\<forall>h. (P (h # t)))))"
+ − 1204
shows "(\<forall>l. (P l))"
+ − 1205
sorry
+ − 1206
+ − 1207
ML {* atomize_thm @{thm list_induct_hol4} *}
99
+ − 1208
+ − 1209
(*fun prove_reg trm \<Rightarrow> thm (we might need some facts to do this)
+ − 1210
trm == new_trm
+ − 1211
*)
+ − 1212
114
+ − 1213
ML {* val li = Thm.freezeT (atomize_thm @{thm list_induct_hol4}) *}
87
+ − 1214
107
+ − 1215
prove list_induct_r: {*
93
+ − 1216
Logic.mk_implies
94
+ − 1217
((prop_of li),
+ − 1218
(regularise (prop_of li) @{typ "'a List.list"} @{term "op \<approx>"} @{context})) *}
93
+ − 1219
apply (simp only: equiv_res_forall[OF equiv_list_eq])
94
+ − 1220
thm RIGHT_RES_FORALL_REGULAR
+ − 1221
apply (rule RIGHT_RES_FORALL_REGULAR)
93
+ − 1222
prefer 2
+ − 1223
apply (assumption)
104
+ − 1224
apply (metis)
94
+ − 1225
done
87
+ − 1226
107
+ − 1227
ML {* val thm = @{thm list_induct_r} OF [li] *}
129
+ − 1228
ML {* val trm_r = build_goal @{context} thm consts @{typ "'a list"} @{typ "'a fset"} *}
112
+ − 1229
lemmas APPLY_RSP_I = APPLY_RSP[of "(op \<approx> ===> op =) ===> op =" "(ABS_fset ---> id) ---> id" "(REP_fset ---> id) ---> id" "op =" "id" "id"]
114
+ − 1230
lemmas REP_ABS_RSP_I = REP_ABS_RSP(1)[of "(op \<approx> ===> op =) ===> op =" "(ABS_fset ---> id) ---> id" "(REP_fset ---> id) ---> id"]
136
+ − 1231
lemmas APPLY_RSP_II = APPLY_RSP[of "op \<approx>" "ABS_fset" "REP_fset" "op \<approx>" "ABS_fset" "REP_fset"]
112
+ − 1232
129
+ − 1233
ML {* val trm = build_goal_term @{context} thm consts @{typ "'a list"} @{typ "'a fset"} *}
+ − 1234
136
+ − 1235
129
+ − 1236
prove list_induct_tr: trm_r
112
+ − 1237
apply (atomize(full))
113
+ − 1238
apply (simp only: id_def[symmetric])
112
+ − 1239
114
+ − 1240
(* APPLY_RSP_TAC *)
112
+ − 1241
ML_prf {*
114
+ − 1242
val (_, gc') = Thm.dest_comb (Subgoal.focus @{context} 1 (Isar.goal ()) |> fst |> #concl)
+ − 1243
val (_, tc') = Thm.dest_comb (cterm_of @{theory} (concl_of @{thm "APPLY_RSP_I" }))
+ − 1244
val m = Thm.match (tc', gc')
+ − 1245
val t2 = Drule.instantiate m @{thm "APPLY_RSP_I" }
112
+ − 1246
*}
126
+ − 1247
thm APPLY_RSP_I
112
+ − 1248
apply (tactic {* rtac t2 1 *})
126
+ − 1249
prefer 2
+ − 1250
apply (rule IDENTITY_QUOTIENT)
+ − 1251
prefer 3
112
+ − 1252
(* ABS_REP_RSP_TAC *)
+ − 1253
ML_prf {*
114
+ − 1254
val (_, gc') = Thm.dest_comb (Subgoal.focus @{context} 1 (Isar.goal ()) |> fst |> #concl)
+ − 1255
val (_, tc') = Thm.dest_comb (cterm_of @{theory} (concl_of @{thm "REP_ABS_RSP_I" }))
+ − 1256
val m = Thm.match (tc', gc')
+ − 1257
val t2 = Drule.instantiate m @{thm "REP_ABS_RSP_I" }
112
+ − 1258
*}
114
+ − 1259
apply (tactic {* rtac t2 1 *})
+ − 1260
prefer 2
+ − 1261
+ − 1262
(* LAMBDA_RES_TAC *)
+ − 1263
apply (simp only: FUN_REL.simps)
+ − 1264
apply (rule allI)
+ − 1265
apply (rule allI)
+ − 1266
apply (rule impI)
+ − 1267
+ − 1268
(* MK_COMB_TAC *)
+ − 1269
apply (tactic {* Cong_Tac.cong_tac @{thm cong} 1 *})
+ − 1270
(* MK_COMB_TAC *)
+ − 1271
apply (tactic {* Cong_Tac.cong_tac @{thm cong} 1 *})
+ − 1272
(* REFL_TAC *)
+ − 1273
apply (simp)
+ − 1274
(* MK_COMB_TAC *)
+ − 1275
apply (tactic {* Cong_Tac.cong_tac @{thm cong} 1 *})
+ − 1276
(* MK_COMB_TAC *)
+ − 1277
apply (tactic {* Cong_Tac.cong_tac @{thm cong} 1 *})
+ − 1278
(* REFL_TAC *)
+ − 1279
apply (simp)
+ − 1280
(* APPLY_RSP_TAC *)
121
+ − 1281
apply (rule_tac APPLY_RSP[of "op \<approx>" "ABS_fset" "REP_fset" "op =" "id" "id"])
120
+ − 1282
apply (rule QUOTIENT_fset)
126
+ − 1283
apply (rule IDENTITY_QUOTIENT)
+ − 1284
prefer 2
120
+ − 1285
(* ABS_REP_RSP *)
+ − 1286
apply (rule REP_ABS_RSP(1)[of "op \<approx>" "ABS_fset" "REP_fset"])
+ − 1287
apply (rule QUOTIENT_fset)
+ − 1288
(* MINE *)
+ − 1289
apply (rule list_eq_refl )
+ − 1290
(* ABS_REP_RSP *)
+ − 1291
apply (rule REP_ABS_RSP(1)[of "op \<approx> ===> op =" "REP_fset ---> id" "ABS_fset ---> id"])
+ − 1292
prefer 2
+ − 1293
(* MINE *)
+ − 1294
apply (simp only: FUN_REL.simps)
126
+ − 1295
prefer 2
120
+ − 1296
(* APPLY_RSP *)
124
+ − 1297
apply (rule_tac APPLY_RSP[of "op \<approx> ===> op =" "REP_fset ---> id" "ABS_fset ---> id" "op =" "id" "id" "Ball (Respects op \<approx>)" "Ball (Respects op \<approx>)"])
126
+ − 1298
prefer 2
+ − 1299
apply (rule IDENTITY_QUOTIENT)
+ − 1300
(* 2: ho_respects *)
+ − 1301
prefer 3
124
+ − 1302
(* ABS_REP_RSP *)
+ − 1303
apply (rule REP_ABS_RSP(1)[of "op \<approx> ===> op =" "REP_fset ---> id" "ABS_fset ---> id"])
+ − 1304
prefer 2
+ − 1305
(* LAMBDA_RSP *)
+ − 1306
apply (simp only: FUN_REL.simps)
+ − 1307
apply (rule allI)
+ − 1308
apply (rule allI)
+ − 1309
apply (rule impI)
+ − 1310
(* MK_COMB_TAC *)
+ − 1311
apply (tactic {* Cong_Tac.cong_tac @{thm cong} 1 *})
+ − 1312
(* MK_COMB_TAC *)
+ − 1313
apply (tactic {* Cong_Tac.cong_tac @{thm cong} 1 *})
+ − 1314
(* REFL_TAC *)
+ − 1315
apply (simp)
+ − 1316
(* APPLY_RSP *)
+ − 1317
apply (rule APPLY_RSP[of "op \<approx>" "ABS_fset" "REP_fset" "op =" "id" "id"])
126
+ − 1318
apply (rule QUOTIENT_fset)
+ − 1319
apply (rule IDENTITY_QUOTIENT)
124
+ − 1320
apply (rule REP_ABS_RSP(1)[of "op \<approx> ===> op =" "REP_fset ---> id" "ABS_fset ---> id"])
+ − 1321
prefer 2
123
+ − 1322
(* MINE *)
124
+ − 1323
apply (simp only: FUN_REL.simps)
126
+ − 1324
prefer 2
124
+ − 1325
apply (rule REP_ABS_RSP(1)[of "op \<approx>" "ABS_fset" "REP_fset"])
126
+ − 1326
apply (rule QUOTIENT_fset)
124
+ − 1327
(* FIRST_ASSUM MATCH_ACCEPT_TAC *)
+ − 1328
apply (assumption)
126
+ − 1329
prefer 2
124
+ − 1330
(* MK_COMB_TAC *)
+ − 1331
apply (tactic {* Cong_Tac.cong_tac @{thm cong} 1 *})
+ − 1332
(* REFL_TAC *)
+ − 1333
apply (simp)
+ − 1334
(* W(C (curry op THEN) (G... *)
+ − 1335
apply (rule ext)
+ − 1336
(* APPLY_RSP *)
+ − 1337
apply (rule APPLY_RSP[of "op \<approx>" "ABS_fset" "REP_fset" "op =" "id" "id"])
126
+ − 1338
apply (rule QUOTIENT_fset)
+ − 1339
apply (rule IDENTITY_QUOTIENT)
124
+ − 1340
apply (rule REP_ABS_RSP(1)[of "op \<approx> ===> op =" "REP_fset ---> id" "ABS_fset ---> id"])
+ − 1341
prefer 2
+ − 1342
apply (simp only: FUN_REL.simps)
126
+ − 1343
prefer 2
124
+ − 1344
apply (rule REP_ABS_RSP(1)[of "op \<approx>" "ABS_fset" "REP_fset"])
126
+ − 1345
apply (rule QUOTIENT_fset)
124
+ − 1346
(* APPLY_RSP *)
136
+ − 1347
+ − 1348
+ − 1349
+ − 1350
ML_prf {*
+ − 1351
val goal = hd (prems_of (Isar.goal ()));
+ − 1352
val goalc = Logic.strip_assums_concl goal
138
+ − 1353
val ps = rev (map Free (Logic.strip_params goal));
+ − 1354
val (_, gc') = Thm.dest_comb (cterm_of @{theory} (Term.subst_bounds (ps, goalc)));
136
+ − 1355
val (_, tc') = Thm.dest_comb (cterm_of @{theory} (concl_of @{thm "APPLY_RSP_II" }))
+ − 1356
*}
138
+ − 1357
ML_prf fold
136
+ − 1358
ML_prf {*
+ − 1359
val m = Thm.match (tc', gc')
138
+ − 1360
val l = map (fst) (snd m);
+ − 1361
val f_inst = map (term_of o snd) (snd m);
+ − 1362
val f_abs = map (fold lambda ps) f_inst
136
+ − 1363
*}
138
+ − 1364
ML_prf {* val f_abs_c = map (cterm_of @{theory}) f_abs *}
+ − 1365
ML_prf {* val insts = l ~~ f_abs_c *}
+ − 1366
ML_prf {* val t = Thm.lift_rule (hd (cprems_of (Isar.goal ()))) @{thm "APPLY_RSP_II" } *}
+ − 1367
ML_prf {* Toplevel.program (fn () => cterm_instantiate insts t) *}
+ − 1368
+ − 1369
ML_prf {* Toplevel.program (fn () => Drule.instantiate' [] [SOME f_abs_c, SOME f_abs_c] t) *}
+ − 1370
+ − 1371
+ − 1372
+ − 1373
.
136
+ − 1374
(* FAILS: t2 = Drule.instantiate m @{thm "APPLY_RSP_II" }; rtac t2 1 *)
137
+ − 1375
(* Works but why does it take a string? *)
+ − 1376
ML_prf {*
+ − 1377
val t_inst = snd m;
+ − 1378
val t_insts = map (fn (a, b) => (fst (dest_Var (term_of a)), (Syntax.string_of_term @{context} (term_of b)))) t_inst
+ − 1379
+ − 1380
*}
+ − 1381
ML_prf dest_Var
+ − 1382
apply (tactic {* res_inst_tac @{context} [(("f",0),"op # h"),(("g",0),"op # h")] @{thm "APPLY_RSP_II" } 1 *})
+ − 1383
ML_prf induct_tac
136
+ − 1384
(*apply (rule_tac ?f="\<lambda>x. h # x" and ?g="\<lambda>x. h # x" in APPLY_RSP[of "op \<approx>" "ABS_fset" "REP_fset" "op \<approx>" "ABS_fset" "REP_fset"] )*)
+ − 1385
126
+ − 1386
apply (rule QUOTIENT_fset)
+ − 1387
apply (rule QUOTIENT_fset)
124
+ − 1388
apply (rule_tac ?f="op #" and ?g="op #" in APPLY_RSP[of "op =" "id" "id" "op \<approx> ===> op \<approx>" "REP_fset ---> ABS_fset" "ABS_fset ---> REP_fset"])
126
+ − 1389
apply (rule IDENTITY_QUOTIENT)
+ − 1390
(* CONS respects *)
+ − 1391
prefer 2
+ − 1392
apply (simp add: FUN_REL.simps)
+ − 1393
apply (rule allI)
124
+ − 1394
apply (rule allI)
+ − 1395
apply (rule allI)
+ − 1396
apply (rule impI)
+ − 1397
apply (rule cons_preserves)
+ − 1398
apply (assumption)
126
+ − 1399
prefer 2
124
+ − 1400
apply (simp)
111
+ − 1401
sorry
107
+ − 1402
129
+ − 1403
prove list_induct_t: trm
+ − 1404
apply (simp only: list_induct_tr[symmetric])
+ − 1405
apply (tactic {* rtac thm 1 *})
+ − 1406
done
+ − 1407
+ − 1408
ML {* val nthm = MetaSimplifier.rewrite_rule fset_defs_sym (snd (no_vars (Context.Theory @{theory}, @{thm list_induct_t}))) *}
+ − 1409
112
+ − 1410
thm list.recs(2)
+ − 1411
+ − 1412
106
+ − 1413
ML {* val card1_suc_f = Thm.freezeT (atomize_thm @{thm card1_suc}) *}
96
4da714704611
A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1414
4da714704611
A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1415
prove card1_suc_r: {*
4da714704611
A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1416
Logic.mk_implies
106
+ − 1417
((prop_of card1_suc_f),
+ − 1418
(regularise (prop_of card1_suc_f) @{typ "'a List.list"} @{term "op \<approx>"} @{context})) *}
104
+ − 1419
apply (simp add: equiv_res_forall[OF equiv_list_eq] equiv_res_exists[OF equiv_list_eq])
96
4da714704611
A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1420
done
4da714704611
A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1421
106
+ − 1422
ML {* @{thm card1_suc_r} OF [card1_suc_f] *}
96
4da714704611
A number of lemmas for REGULARIZE_TAC and regularizing card1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1423
97
+ − 1424
ML {* val li = Thm.freezeT (atomize_thm @{thm fold1.simps(2)}) *}
+ − 1425
prove fold1_def_2_r: {*
+ − 1426
Logic.mk_implies
+ − 1427
((prop_of li),
+ − 1428
(regularise (prop_of li) @{typ "'a List.list"} @{term "op \<approx>"} @{context})) *}
+ − 1429
apply (simp add: equiv_res_forall[OF equiv_list_eq])
+ − 1430
done
+ − 1431
+ − 1432
ML {* @{thm fold1_def_2_r} OF [li] *}
+ − 1433
29
+ − 1434
ML {*
49
+ − 1435
fun lift_theorem_fset_aux thm lthy =
+ − 1436
let
+ − 1437
val ((_, [novars]), lthy2) = Variable.import true [thm] lthy;
111
+ − 1438
val goal = build_goal @{context} novars consts @{typ "'a list"} @{typ "'a fset"};
49
+ − 1439
val cgoal = cterm_of @{theory} goal;
+ − 1440
val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite true fset_defs_sym cgoal);
106
+ − 1441
val tac = transconv_fset_tac' @{context};
49
+ − 1442
val cthm = Goal.prove_internal [] cgoal2 (fn _ => tac);
+ − 1443
val nthm = MetaSimplifier.rewrite_rule [symmetric cthm] (snd (no_vars (Context.Theory @{theory}, thm)))
+ − 1444
val nthm2 = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.REPS_same QUOT_TYPE_I_fset.thm10} nthm;
+ − 1445
val [nthm3] = ProofContext.export lthy2 lthy [nthm2]
+ − 1446
in
+ − 1447
nthm3
+ − 1448
end
+ − 1449
*}
+ − 1450
+ − 1451
ML {* lift_theorem_fset_aux @{thm m1} @{context} *}
+ − 1452
+ − 1453
ML {*
+ − 1454
fun lift_theorem_fset name thm lthy =
+ − 1455
let
+ − 1456
val lifted_thm = lift_theorem_fset_aux thm lthy;
88
+ − 1457
val (_, lthy2) = note (name, lifted_thm) lthy;
49
+ − 1458
in
+ − 1459
lthy2
+ − 1460
end;
+ − 1461
*}
+ − 1462
+ − 1463
local_setup {* lift_theorem_fset @{binding "m1_lift"} @{thm m1} *}
+ − 1464
local_setup {* lift_theorem_fset @{binding "leqi4_lift"} @{thm list_eq.intros(4)} *}
+ − 1465
local_setup {* lift_theorem_fset @{binding "leqi5_lift"} @{thm list_eq.intros(5)} *}
+ − 1466
local_setup {* lift_theorem_fset @{binding "m2_lift"} @{thm m2} *}
+ − 1467
thm m1_lift
+ − 1468
thm leqi4_lift
+ − 1469
thm leqi5_lift
+ − 1470
thm m2_lift
106
+ − 1471
ML {* @{thm card1_suc_r} OF [card1_suc_f] *}
+ − 1472
(*ML {* Toplevel.program (fn () => lift_theorem_fset @{binding "card_suc"}
+ − 1473
(@{thm card1_suc_r} OF [card1_suc_f]) @{context}) *}*)
+ − 1474
(*local_setup {* lift_theorem_fset @{binding "card_suc"} @{thm card1_suc} *}*)
49
+ − 1475
+ − 1476
thm leqi4_lift
+ − 1477
ML {*
62
+ − 1478
val (nam, typ) = hd (Term.add_vars (prop_of @{thm leqi4_lift}) [])
49
+ − 1479
val (_, l) = dest_Type typ
+ − 1480
val t = Type ("QuotMain.fset", l)
+ − 1481
val v = Var (nam, t)
+ − 1482
val cv = cterm_of @{theory} ((term_of @{cpat "REP_fset"}) $ v)
+ − 1483
*}
+ − 1484
+ − 1485
ML {*
+ − 1486
Toplevel.program (fn () =>
+ − 1487
MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.thm10} (
+ − 1488
Drule.instantiate' [] [NONE, SOME (cv)] @{thm leqi4_lift}
+ − 1489
)
+ − 1490
)
+ − 1491
*}
+ − 1492
77
+ − 1493
lemma "(Ball (Respects ((op \<approx>) ===> (op =)))
72
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1494
(((REP_fset ---> id) ---> id)
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1495
(((ABS_fset ---> id) ---> id)
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1496
(\<lambda>P.
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1497
(ABS_fset ---> id) ((REP_fset ---> id) P)
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1498
(REP_fset (ABS_fset [])) \<and>
77
+ − 1499
Ball (Respects (op \<approx>))
72
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1500
((ABS_fset ---> id)
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1501
((REP_fset ---> id)
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1502
(\<lambda>t.
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1503
((ABS_fset ---> id)
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1504
((REP_fset ---> id) P)
73
+ − 1505
(REP_fset (ABS_fset t))) \<longrightarrow>
72
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1506
(!h.
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1507
(ABS_fset ---> id)
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1508
((REP_fset ---> id) P)
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1509
(REP_fset
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1510
(ABS_fset
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1511
(h #
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1512
REP_fset
73
+ − 1513
(ABS_fset t)))))))) \<longrightarrow>
77
+ − 1514
Ball (Respects (op \<approx>))
72
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1515
((ABS_fset ---> id)
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1516
((REP_fset ---> id)
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1517
(\<lambda>l.
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1518
(ABS_fset ---> id)
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1519
((REP_fset ---> id) P)
74
+ − 1520
(REP_fset (ABS_fset l)))))))))
+ − 1521
77
+ − 1522
= Ball (Respects ((op \<approx>) ===> (op =)))
+ − 1523
(\<lambda>P. ((P []) \<and> (Ball (Respects (op \<approx>)) (\<lambda>t. P t \<longrightarrow> (\<forall>h. (P (h # t)))))) \<longrightarrow>
+ − 1524
(Ball (Respects (op \<approx>)) (\<lambda>l. P l)))"
112
+ − 1525
+ − 1526
thm APPLY_RSP[of "op \<approx> ===> (op = ===> op =)" "REP_fset ---> (id ---> id)" "ABS_fset ---> (id ---> id)" "op =" "id" "id"]
+ − 1527
proof -
+ − 1528
note APPLY_RSP_I = APPLY_RSP[of "op \<approx> ===> (op = ===> op =)" "REP_fset ---> (id ---> id)" "ABS_fset ---> (id ---> id)" "op =" "id" "id"]
+ − 1529
show ?thesis apply -
+ − 1530
+ − 1531
+ − 1532
ML_prf {*
+ − 1533
val gc = Subgoal.focus @{context} 1 (Isar.goal ())
+ − 1534
|> fst
+ − 1535
|> #concl
+ − 1536
val tc = cterm_of @{theory} (concl_of ( @{thm "APPLY_RSP" }))
+ − 1537
val (_, tc') = Thm.dest_comb tc
+ − 1538
val (_, gc') = Thm.dest_comb gc
+ − 1539
*}
+ − 1540
+ − 1541
+ − 1542
+ − 1543
+ − 1544
ML_prf {*
+ − 1545
(gc', tc')
+ − 1546
*}
+ − 1547
ML_prf {* Pattern.first_order_match *}
+ − 1548
ML_prf {* Thm.match (@{cpat "?P::?'a"}, @{cterm "1::nat"}) *}
+ − 1549
ML_prf {* Pattern.match @{theory} (term_of @{cpat "?P::nat"}, term_of @{cterm "1::nat"}) (Vartab.empty, Vartab.empty) *}
+ − 1550
+ − 1551
ML_prf {* val (op1, (l1, r1)) = dest_cbinop tc' *}
+ − 1552
ML_prf {* val (op2, (l2, r2)) = dest_cbinop gc' *}
+ − 1553
+ − 1554
ML_prf {* val mo = Thm.match (op1, op2) *}
+ − 1555
ML_prf {* val t1 = Drule.instantiate mo @{thm "APPLY_RSP" } *}
+ − 1556
ML_prf {* val tc = cterm_of @{theory} (concl_of t1) *}
+ − 1557
ML_prf {* val (_, tc') = Thm.dest_comb tc *}
+ − 1558
ML_prf {* val (op1, (l1, r1)) = dest_cbinop tc' *}
+ − 1559
ML_prf {* val (mlty, mlt) = Thm.match (l1, l2) *}
+ − 1560
ML_prf {* val t2 = Drule.instantiate (mlty, []) t1 *}
+ − 1561
ML_prf {* val tc = cterm_of @{theory} (concl_of t2) *}
+ − 1562
ML_prf {* val (_, tc') = Thm.dest_comb tc *}
+ − 1563
ML_prf {* val (op1, (l1, r1)) = dest_cbinop tc' *}
+ − 1564
ML_prf {* val mr = Thm.match (r1, r2) *}
+ − 1565
+ − 1566
ML_prf {* val t3 = Drule.instantiate ml t2 *}
+ − 1567
ML_prf {* val tc = cterm_of @{theory} (concl_of t3) *}
+ − 1568
ML_prf {* val (_, tc') = Thm.dest_comb tc *}
+ − 1569
+ − 1570
+ − 1571
+ − 1572
ML_prf {* mtyl; mtyr *}
+ − 1573
+ − 1574
ML_prf {* val ol1 = Thm.capply op1 l1 *}
+ − 1575
ML_prf {* val ol2 = Thm.capply op2 l2 *}
+ − 1576
ML_prf {* (ol1, ol2); Thm.match (ol1, ol2) *}
+ − 1577
ML_prf {* val w1 = Thm.capply ol1 r1 *}
+ − 1578
ML_prf {* val w2 = Thm.capply ol2 r2 *}
+ − 1579
+ − 1580
+ − 1581
.
+ − 1582
(*ML_prf {* (w1, w2); Thm.match (w1, w2) *}*)
+ − 1583
+ − 1584
+ − 1585
+ − 1586
ML_prf {*
+ − 1587
(tc', gc')
+ − 1588
*}
+ − 1589
95
+ − 1590
thm APPLY_RSP
112
+ − 1591
thm APPLY_RSP[of "op \<approx> ===> (op = ===> op =)"]
+ − 1592
term "(op \<approx> ===> op =) ===> op ="
+ − 1593
proof -
+ − 1594
note APPLY_RSP_I2 = APPLY_RSP[of "(op \<approx> ===> op =) ===> op =" "(ABS_fset ---> id) ---> id" "(REP_fset ---> id) ---> id" "op =" "id" "id" "Ball (Respects op \<approx> ===> op =)" "Ball (Respects op \<approx> ===> op =)"]
+ − 1595
show ?thesis apply -
+ − 1596
+ − 1597
thm APPLY_RSP_I2[of _ "\<lambda>P. (P [] \<and> (\<forall>t\<Colon>'b list\<in>Respects op \<approx>. P t \<longrightarrow> (\<forall>h\<Colon>'b. P (h # t))) \<longrightarrow>
+ − 1598
(\<forall>l\<Colon>'b list\<in>Respects op \<approx>. P l))"]
+ − 1599
proof -
+ − 1600
note APPLY_RSP_I3=APPLY_RSP_I2[of _ "\<lambda>P. (P [] \<and> (\<forall>t\<Colon>'b list\<in>Respects op \<approx>. P t \<longrightarrow> (\<forall>h\<Colon>'b. P (h # t))) \<longrightarrow>
+ − 1601
(\<forall>l\<Colon>'b list\<in>Respects op \<approx>. P l))"]
+ − 1602
show ?thesis apply -
+ − 1603
term "(REP_fset ---> id ---> id
+ − 1604
(ABS_fset ---> id ---> id
+ − 1605
(\<lambda>P\<Colon>'a list \<Rightarrow> bool.
+ − 1606
ABS_fset ---> id (REP_fset ---> id P)
+ − 1607
(REP_fset (ABS_fset [])) \<and>
+ − 1608
Ball (Respects op \<approx>)
+ − 1609
(ABS_fset ---> id
+ − 1610
(REP_fset ---> id
+ − 1611
(\<lambda>t\<Colon>'a list.
+ − 1612
ABS_fset ---> id (REP_fset ---> id P)
+ − 1613
(REP_fset (ABS_fset t)) \<longrightarrow>
+ − 1614
(\<forall>h\<Colon>'a.
+ − 1615
ABS_fset ---> id (REP_fset ---> id P)
+ − 1616
(REP_fset
+ − 1617
(ABS_fset
+ − 1618
(h # REP_fset (ABS_fset t)))))))) \<longrightarrow>
+ − 1619
Ball (Respects op \<approx>)
+ − 1620
(ABS_fset ---> id
+ − 1621
(REP_fset ---> id
+ − 1622
(\<lambda>l\<Colon>'a list.
+ − 1623
ABS_fset ---> id (REP_fset ---> id P)
+ − 1624
(REP_fset (ABS_fset l))))))))"
+ − 1625
+ − 1626
apply (rule APPLY_RSP_I3)
+ − 1627
term "Ball (Respects op \<approx> ===> op =)"
+ − 1628
thm APPLY_RSP_I [of _ "Ball (Respects op \<approx> ===> op =)"]
+ − 1629
+ − 1630
thm APPLY_RSP
+ − 1631
+ − 1632
.
+ − 1633
+ − 1634
apply (tactic {* Cong_Tac.cong_tac @{thm APPLY_RSP_I} 1 *} )
+ − 1635
+ − 1636
apply (tactic {* match_tac @{thms APPLY_RSP_I} 1 *})
+ − 1637
.
+ − 1638
111
+ − 1639
apply (rule APPLY_RSP[of "op \<approx> ===> (op = ===> op =)" "REP_fset ---> (id ---> id)" "ABS_fset ---> (id ---> id)" "op =" "id" "id"])
95
+ − 1640
apply (rule APPLY_RSP[of "op \<approx> ===> (op = ===> op =)" _ _ "op ="])
74
+ − 1641
term "(\<forall>P. (((P []) \<and> (\<forall>t. (P t \<longrightarrow> (\<forall>h. P (h # t))))) \<longrightarrow> (\<forall>l. (P l))))"
73
+ − 1642
thm LAMBDA_PRS1[symmetric]
+ − 1643
(*apply (simp only:LAMBDA_PRS1[symmetric] FUN_QUOTIENT IDENTITY_QUOTIENT QUOT_TYPE_I_fset.QUOTIENT)*)
77
+ − 1644
apply (unfold Ball_def)
74
+ − 1645
apply (simp only: IN_RESPECTS)
+ − 1646
apply (simp only:list_eq_refl)
72
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1647
apply (unfold id_def)
74
+ − 1648
(*apply (simp only: FUN_MAP_I)*)
+ − 1649
apply (simp)
+ − 1650
apply (simp only: QUOT_TYPE_I_fset.thm10)
+ − 1651
..
73
+ − 1652
apply (simp add:IN_RESPECTS)
74
+ − 1653
73
+ − 1654
+ − 1655
+ − 1656
+ − 1657
apply (simp add: QUOT_TYPE_I_fset.R_trans2)
+ − 1658
+ − 1659
apply (rule ext)
+ − 1660
apply (simp add:QUOT_TYPE_I_fset.REP_ABS_rsp)
+ − 1661
apply (tactic {* Cong_Tac.cong_tac @{thm cong} 1 *} )
+ − 1662
apply (simp add:cons_preserves)
+ − 1663
+ − 1664
72
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1665
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1666
(*prove aaa: {* (Thm.term_of cgoal2) *}
45
+ − 1667
apply (tactic {* LocalDefs.unfold_tac @{context} fset_defs *} )
58
f2565006dc5a
Just one atomize is enough for the currently lifted theorems. Properly lift 'all' and 'Ex'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1668
apply (atomize(full))
101
+ − 1669
apply (tactic {* transconv_fset_tac' @{context} 1 *})
72
4efc9e6661a4
Testing if I can prove the regularized version of induction manually
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1670
done*)
29
+ − 1671
18
+ − 1672
(*
+ − 1673
datatype obj1 =
+ − 1674
OVAR1 "string"
+ − 1675
| OBJ1 "(string * (string \<Rightarrow> obj1)) list"
+ − 1676
| INVOKE1 "obj1 \<Rightarrow> string"
+ − 1677
| UPDATE1 "obj1 \<Rightarrow> string \<Rightarrow> (string \<Rightarrow> obj1)"
+ − 1678
*)
53
+ − 1679
57
+ − 1680
end
112
+ − 1681