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