163
+ − 1
theory FSet
+ − 2
imports QuotMain
+ − 3
begin
+ − 4
+ − 5
inductive
+ − 6
list_eq (infix "\<approx>" 50)
+ − 7
where
+ − 8
"a#b#xs \<approx> b#a#xs"
+ − 9
| "[] \<approx> []"
+ − 10
| "xs \<approx> ys \<Longrightarrow> ys \<approx> xs"
+ − 11
| "a#a#xs \<approx> a#xs"
+ − 12
| "xs \<approx> ys \<Longrightarrow> a#xs \<approx> a#ys"
+ − 13
| "\<lbrakk>xs1 \<approx> xs2; xs2 \<approx> xs3\<rbrakk> \<Longrightarrow> xs1 \<approx> xs3"
+ − 14
+ − 15
lemma list_eq_refl:
+ − 16
shows "xs \<approx> xs"
+ − 17
apply (induct xs)
+ − 18
apply (auto intro: list_eq.intros)
+ − 19
done
+ − 20
+ − 21
lemma equiv_list_eq:
+ − 22
shows "EQUIV list_eq"
+ − 23
unfolding EQUIV_REFL_SYM_TRANS REFL_def SYM_def TRANS_def
+ − 24
apply(auto intro: list_eq.intros list_eq_refl)
+ − 25
done
+ − 26
+ − 27
quotient fset = "'a list" / "list_eq"
+ − 28
apply(rule equiv_list_eq)
+ − 29
done
+ − 30
+ − 31
print_theorems
+ − 32
+ − 33
typ "'a fset"
+ − 34
thm "Rep_fset"
+ − 35
thm "ABS_fset_def"
+ − 36
268
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 37
quotient_def
231
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 38
EMPTY :: "'a fset"
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 39
where
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 40
"EMPTY \<equiv> ([]::'a list)"
163
+ − 41
+ − 42
term Nil
+ − 43
term EMPTY
+ − 44
thm EMPTY_def
+ − 45
268
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 46
quotient_def
254
+ − 47
INSERT :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
231
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 48
where
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 49
"INSERT \<equiv> op #"
163
+ − 50
+ − 51
term Cons
+ − 52
term INSERT
+ − 53
thm INSERT_def
+ − 54
268
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 55
quotient_def
231
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 56
FUNION :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 57
where
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 58
"FUNION \<equiv> (op @)"
163
+ − 59
+ − 60
term append
231
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 61
term FUNION
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 62
thm FUNION_def
163
+ − 63
+ − 64
thm QUOTIENT_fset
+ − 65
+ − 66
thm QUOT_TYPE_I_fset.thm11
+ − 67
+ − 68
+ − 69
fun
+ − 70
membship :: "'a \<Rightarrow> 'a list \<Rightarrow> bool" (infix "memb" 100)
+ − 71
where
+ − 72
m1: "(x memb []) = False"
+ − 73
| m2: "(x memb (y#xs)) = ((x=y) \<or> (x memb xs))"
+ − 74
+ − 75
fun
+ − 76
card1 :: "'a list \<Rightarrow> nat"
+ − 77
where
+ − 78
card1_nil: "(card1 []) = 0"
+ − 79
| card1_cons: "(card1 (x # xs)) = (if (x memb xs) then (card1 xs) else (Suc (card1 xs)))"
+ − 80
268
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 81
quotient_def
231
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 82
CARD :: "'a fset \<Rightarrow> nat"
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 83
where
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 84
"CARD \<equiv> card1"
163
+ − 85
+ − 86
term card1
231
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 87
term CARD
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 88
thm CARD_def
163
+ − 89
+ − 90
(* text {*
+ − 91
Maybe make_const_def should require a theorem that says that the particular lifted function
+ − 92
respects the relation. With it such a definition would be impossible:
+ − 93
make_const_def @{binding CARD} @{term "length"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
+ − 94
*}*)
+ − 95
+ − 96
lemma card1_0:
+ − 97
fixes a :: "'a list"
+ − 98
shows "(card1 a = 0) = (a = [])"
214
+ − 99
by (induct a) auto
163
+ − 100
+ − 101
lemma not_mem_card1:
+ − 102
fixes x :: "'a"
+ − 103
fixes xs :: "'a list"
+ − 104
shows "~(x memb xs) \<Longrightarrow> card1 (x # xs) = Suc (card1 xs)"
+ − 105
by simp
+ − 106
+ − 107
lemma mem_cons:
+ − 108
fixes x :: "'a"
+ − 109
fixes xs :: "'a list"
+ − 110
assumes a : "x memb xs"
+ − 111
shows "x # xs \<approx> xs"
214
+ − 112
using a by (induct xs) (auto intro: list_eq.intros )
163
+ − 113
+ − 114
lemma card1_suc:
+ − 115
fixes xs :: "'a list"
+ − 116
fixes n :: "nat"
+ − 117
assumes c: "card1 xs = Suc n"
+ − 118
shows "\<exists>a ys. ~(a memb ys) \<and> xs \<approx> (a # ys)"
+ − 119
using c
+ − 120
apply(induct xs)
+ − 121
apply (metis Suc_neq_Zero card1_0)
+ − 122
apply (metis QUOT_TYPE_I_fset.R_trans card1_cons list_eq_refl mem_cons)
+ − 123
done
+ − 124
+ − 125
primrec
+ − 126
fold1
+ − 127
where
+ − 128
"fold1 f (g :: 'a \<Rightarrow> 'b) (z :: 'b) [] = z"
+ − 129
| "fold1 f g z (a # A) =
+ − 130
(if ((!u v. (f u v = f v u))
+ − 131
\<and> (!u v w. ((f u (f v w) = f (f u v) w))))
+ − 132
then (
+ − 133
if (a memb A) then (fold1 f g z A) else (f (g a) (fold1 f g z A))
+ − 134
) else z)"
+ − 135
+ − 136
(* fold1_def is not usable, but: *)
+ − 137
thm fold1.simps
+ − 138
+ − 139
lemma fs1_strong_cases:
+ − 140
fixes X :: "'a list"
+ − 141
shows "(X = []) \<or> (\<exists>a. \<exists> Y. (~(a memb Y) \<and> (X \<approx> a # Y)))"
+ − 142
apply (induct X)
+ − 143
apply (simp)
+ − 144
apply (metis QUOT_TYPE_I_fset.thm11 list_eq_refl mem_cons m1)
+ − 145
done
+ − 146
268
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 147
quotient_def
231
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 148
IN :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool"
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 149
where
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 150
"IN \<equiv> membship"
163
+ − 151
+ − 152
term membship
+ − 153
term IN
+ − 154
thm IN_def
+ − 155
274
+ − 156
term fold1
+ − 157
quotient_def
+ − 158
FOLD :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b fset \<Rightarrow> 'a"
231
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 159
where
c643938b846a
updated some definitions; had to give sometimes different names; somewhere I introduced a bug, since not everything is working anymore (needs fixing!)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 160
"FOLD \<equiv> fold1"
194
+ − 161
+ − 162
term fold1
+ − 163
term fold
+ − 164
thm fold_def
+ − 165
268
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 166
quotient_def
254
+ − 167
fmap::"('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset"
225
+ − 168
where
254
+ − 169
"fmap \<equiv> map"
194
+ − 170
+ − 171
term map
+ − 172
term fmap
+ − 173
thm fmap_def
+ − 174
274
+ − 175
ML {* val defs = @{thms EMPTY_def IN_def FUNION_def CARD_def INSERT_def fmap_def FOLD_def} *}
+ − 176
ML {* val consts = lookup_quot_consts defs *}
+ − 177
ML {* val defs_sym = add_lower_defs @{context} defs *}
163
+ − 178
164
+ − 179
lemma memb_rsp:
163
+ − 180
fixes z
+ − 181
assumes a: "list_eq x y"
+ − 182
shows "(z memb x) = (z memb y)"
+ − 183
using a by induct auto
+ − 184
164
+ − 185
lemma ho_memb_rsp:
+ − 186
"(op = ===> (op \<approx> ===> op =)) (op memb) (op memb)"
214
+ − 187
by (simp add: memb_rsp)
164
+ − 188
163
+ − 189
lemma card1_rsp:
+ − 190
fixes a b :: "'a list"
+ − 191
assumes e: "a \<approx> b"
+ − 192
shows "card1 a = card1 b"
214
+ − 193
using e by induct (simp_all add:memb_rsp)
163
+ − 194
228
+ − 195
lemma ho_card1_rsp: "(op \<approx> ===> op =) card1 card1"
214
+ − 196
by (simp add: card1_rsp)
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 197
164
+ − 198
lemma cons_rsp:
163
+ − 199
fixes z
+ − 200
assumes a: "xs \<approx> ys"
+ − 201
shows "(z # xs) \<approx> (z # ys)"
+ − 202
using a by (rule list_eq.intros(5))
+ − 203
164
+ − 204
lemma ho_cons_rsp:
228
+ − 205
"(op = ===> op \<approx> ===> op \<approx>) op # op #"
214
+ − 206
by (simp add: cons_rsp)
164
+ − 207
175
+ − 208
lemma append_rsp_fst:
163
+ − 209
assumes a : "list_eq l1 l2"
214
+ − 210
shows "(l1 @ s) \<approx> (l2 @ s)"
163
+ − 211
using a
214
+ − 212
by (induct) (auto intro: list_eq.intros list_eq_refl)
+ − 213
+ − 214
lemma append_end:
+ − 215
shows "(e # l) \<approx> (l @ [e])"
+ − 216
apply (induct l)
+ − 217
apply (auto intro: list_eq.intros list_eq_refl)
+ − 218
done
+ − 219
+ − 220
lemma rev_rsp:
+ − 221
shows "a \<approx> rev a"
+ − 222
apply (induct a)
+ − 223
apply simp
+ − 224
apply (rule list_eq_refl)
+ − 225
apply simp_all
+ − 226
apply (rule list_eq.intros(6))
+ − 227
prefer 2
+ − 228
apply (rule append_rsp_fst)
+ − 229
apply assumption
+ − 230
apply (rule append_end)
+ − 231
done
163
+ − 232
214
+ − 233
lemma append_sym_rsp:
+ − 234
shows "(a @ b) \<approx> (b @ a)"
+ − 235
apply (rule list_eq.intros(6))
+ − 236
apply (rule append_rsp_fst)
+ − 237
apply (rule rev_rsp)
+ − 238
apply (rule list_eq.intros(6))
+ − 239
apply (rule rev_rsp)
+ − 240
apply (simp)
+ − 241
apply (rule append_rsp_fst)
+ − 242
apply (rule list_eq.intros(3))
+ − 243
apply (rule rev_rsp)
+ − 244
done
+ − 245
+ − 246
lemma append_rsp:
+ − 247
assumes a : "list_eq l1 r1"
+ − 248
assumes b : "list_eq l2 r2 "
+ − 249
shows "(l1 @ l2) \<approx> (r1 @ r2)"
+ − 250
apply (rule list_eq.intros(6))
+ − 251
apply (rule append_rsp_fst)
+ − 252
using a apply (assumption)
+ − 253
apply (rule list_eq.intros(6))
+ − 254
apply (rule append_sym_rsp)
+ − 255
apply (rule list_eq.intros(6))
+ − 256
apply (rule append_rsp_fst)
+ − 257
using b apply (assumption)
+ − 258
apply (rule append_sym_rsp)
+ − 259
done
175
+ − 260
194
+ − 261
lemma ho_append_rsp:
228
+ − 262
"(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"
214
+ − 263
by (simp add: append_rsp)
175
+ − 264
194
+ − 265
lemma map_rsp:
+ − 266
assumes a: "a \<approx> b"
+ − 267
shows "map f a \<approx> map f b"
+ − 268
using a
+ − 269
apply (induct)
+ − 270
apply(auto intro: list_eq.intros)
+ − 271
done
+ − 272
215
+ − 273
lemma fun_rel_id:
228
+ − 274
"(op = ===> op =) \<equiv> op ="
215
+ − 275
apply (rule eq_reflection)
+ − 276
apply (rule ext)
+ − 277
apply (rule ext)
+ − 278
apply (simp)
+ − 279
apply (auto)
+ − 280
apply (rule ext)
+ − 281
apply (simp)
+ − 282
done
+ − 283
194
+ − 284
lemma ho_map_rsp:
228
+ − 285
"((op = ===> op =) ===> op \<approx> ===> op \<approx>) map map"
215
+ − 286
by (simp add: fun_rel_id map_rsp)
194
+ − 287
+ − 288
lemma map_append :
258
+ − 289
"(map f (a @ b)) \<approx>
+ − 290
(map f a) @ (map f b)"
215
+ − 291
by simp (rule list_eq_refl)
194
+ − 292
273
+ − 293
lemma op_eq_twice: "(op = ===> op =) = (op =)"
+ − 294
apply (rule ext)
+ − 295
apply (rule ext)
+ − 296
apply (simp add: FUN_REL.simps)
+ − 297
apply auto
+ − 298
apply (rule ext)
+ − 299
apply simp
+ − 300
done
+ − 301
+ − 302
+ − 303
+ − 304
lemma ho_fold_rsp:
+ − 305
"((op = ===> op = ===> op =) ===> (op = ===> op =) ===> op = ===> op \<approx> ===> op =) fold1 fold1"
+ − 306
apply (auto simp add: op_eq_twice)
+ − 307
sorry
241
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 308
254
+ − 309
print_quotients
+ − 310
+ − 311
226
+ − 312
ML {* val qty = @{typ "'a fset"} *}
+ − 313
ML {* val rsp_thms =
273
+ − 314
@{thms ho_memb_rsp ho_cons_rsp ho_card1_rsp ho_map_rsp ho_append_rsp ho_fold_rsp}
226
+ − 315
@ @{thms ho_all_prs ho_ex_prs} *}
206
+ − 316
239
+ − 317
ML {* fun lift_thm_fset lthy t = lift_thm lthy qty "fset" rsp_thms defs t *}
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 318
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 319
(* ML {* lift_thm_fset @{context} @{thm neq_Nil_conv} *} *)
273
+ − 320
+ − 321
+ − 322
226
+ − 323
ML {* lift_thm_fset @{context} @{thm m1} *}
+ − 324
ML {* lift_thm_fset @{context} @{thm m2} *}
+ − 325
ML {* lift_thm_fset @{context} @{thm list_eq.intros(4)} *}
+ − 326
ML {* lift_thm_fset @{context} @{thm list_eq.intros(5)} *}
+ − 327
ML {* lift_thm_fset @{context} @{thm card1_suc} *}
258
+ − 328
(*ML {* lift_thm_fset @{context} @{thm map_append} *}*)
251
+ − 329
ML {* lift_thm_fset @{context} @{thm append_assoc} *}
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 330
ML {* lift_thm_fset @{context} @{thm list.induct} *}
273
+ − 331
ML {* lift_thm_fset @{context} @{thm fold1.simps(2)} *}
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 332
273
+ − 333
term list_rec
+ − 334
+ − 335
quotient_def
276
+ − 336
fset_rec::"'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
273
+ − 337
where
+ − 338
"fset_rec \<equiv> list_rec"
+ − 339
173
7cf227756e2a
Finally completely lift the previously lifted theorems + clean some old stuff
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 340
thm list.recs(2)
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 341
thm list.cases
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 342
273
+ − 343
+ − 344
+ − 345
+ − 346
+ − 347
+ − 348
ML {* val consts = lookup_quot_consts defs *}
+ − 349
ML {* val defs_sym = add_lower_defs @{context} defs *}
+ − 350
ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
+ − 351
ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "fset" *}
+ − 352
+ − 353
276
+ − 354
ML {* val ind_r_a = atomize_thm @{thm map_append} *}
+ − 355
prove {* build_regularize_goal ind_r_a rty rel @{context} *}
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 356
ML_prf {* fun tac ctxt =
251
+ − 357
(FIRST' [
+ − 358
rtac rel_refl,
+ − 359
atac,
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 360
rtac @{thm get_rid},
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 361
rtac @{thm get_rid2},
251
+ − 362
(fn i => CHANGED (asm_full_simp_tac ((Simplifier.context ctxt HOL_ss) addsimps
239
+ − 363
[(@{thm equiv_res_forall} OF [rel_eqv]),
251
+ − 364
(@{thm equiv_res_exists} OF [rel_eqv])]) i)),
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 365
(rtac @{thm impI} THEN' (asm_full_simp_tac (Simplifier.context ctxt HOL_ss)) THEN' rtac rel_refl),
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 366
(rtac @{thm RIGHT_RES_FORALL_REGULAR})
251
+ − 367
]);
+ − 368
*}
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 369
apply (atomize(full))
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 370
apply (tactic {* tac @{context} 1 *}) *)
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 371
ML {* val ind_r_r = regularize ind_r_a rty rel rel_eqv rel_refl @{context} *}
273
+ − 372
ML {*
239
+ − 373
val rt = build_repabs_term @{context} ind_r_r consts rty qty
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 374
val rg = Logic.mk_equals ((Thm.prop_of ind_r_r), rt);
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 375
*}
257
+ − 376
prove rg
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 377
apply(atomize(full))
257
+ − 378
ML_prf {* fun r_mk_comb_tac_fset lthy = r_mk_comb_tac lthy rty quot rel_refl trans2 rsp_thms *}
194
+ − 379
apply (tactic {* REPEAT_ALL_NEW (r_mk_comb_tac_fset @{context}) 1 *})
273
+ − 380
done
210
+ − 381
ML {* val ind_r_t =
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 382
Toplevel.program (fn () =>
239
+ − 383
repabs @{context} ind_r_r consts rty qty quot rel_refl trans2 rsp_thms
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 384
)
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 385
*}
241
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 386
226
+ − 387
ML {* val abs = findabs rty (prop_of (atomize_thm @{thm list.induct})) *}
241
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 388
ML {* val aps = findaps rty (prop_of (atomize_thm @{thm list.induct})) *}
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 389
ML {* val simp_lam_prs_thms = map (make_simp_prs_thm @{context} quot @{thm LAMBDA_PRS}) abs *}
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 390
ML {* val ind_r_l = repeat_eqsubst_thm @{context} (simp_app_prs_thms @ simp_lam_prs_thms) ind_r_t *}
260
59578f428bbe
Fixes after optimization and preparing for a general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 391
ML {* val thm = @{thm FORALL_PRS[OF FUN_QUOTIENT[OF QUOTIENT_fset IDENTITY_QUOTIENT]]} *}
59578f428bbe
Fixes after optimization and preparing for a general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 392
ML {* val ind_r_a = simp_allex_prs quot [thm] ind_r_l *}
226
+ − 393
ML {* val defs_sym = add_lower_defs @{context} defs *}
260
59578f428bbe
Fixes after optimization and preparing for a general FORALL_PRS
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 394
ML {* val ind_r_d = repeat_eqsubst_thm @{context} defs_sym ind_r_a *}
178
945786a68ec6
Finally lifted induction, with some manually added simplification lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 395
ML {* val ind_r_s = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.REPS_same} ind_r_d *}
209
+ − 396
ML {* ObjectLogic.rulify ind_r_s *}
178
945786a68ec6
Finally lifted induction, with some manually added simplification lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 397
163
+ − 398
ML {*
226
+ − 399
fun lift_thm_fset_note name thm lthy =
163
+ − 400
let
226
+ − 401
val lifted_thm = lift_thm_fset lthy thm;
163
+ − 402
val (_, lthy2) = note (name, lifted_thm) lthy;
+ − 403
in
+ − 404
lthy2
+ − 405
end;
+ − 406
*}
+ − 407
226
+ − 408
local_setup {*
+ − 409
lift_thm_fset_note @{binding "m1l"} @{thm m1} #>
+ − 410
lift_thm_fset_note @{binding "m2l"} @{thm m2} #>
+ − 411
lift_thm_fset_note @{binding "leqi4l"} @{thm list_eq.intros(4)} #>
+ − 412
lift_thm_fset_note @{binding "leqi5l"} @{thm list_eq.intros(5)}
163
+ − 413
*}
226
+ − 414
thm m1l
+ − 415
thm m2l
+ − 416
thm leqi4l
+ − 417
thm leqi5l
163
+ − 418
+ − 419
end