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"
309
+ − 104
shows "(~(x memb xs)) = (card1 (x # xs) = Suc (card1 xs))"
+ − 105
by auto
163
+ − 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
294
+ − 125
definition
+ − 126
rsp_fold
+ − 127
where
+ − 128
"rsp_fold f = ((!u v. (f u v = f v u)) \<and> (!u v w. ((f u (f v w) = f (f u v) w))))"
+ − 129
163
+ − 130
primrec
+ − 131
fold1
+ − 132
where
+ − 133
"fold1 f (g :: 'a \<Rightarrow> 'b) (z :: 'b) [] = z"
+ − 134
| "fold1 f g z (a # A) =
294
+ − 135
(if rsp_fold f
163
+ − 136
then (
+ − 137
if (a memb A) then (fold1 f g z A) else (f (g a) (fold1 f g z A))
+ − 138
) else z)"
+ − 139
+ − 140
(* fold1_def is not usable, but: *)
+ − 141
thm fold1.simps
+ − 142
+ − 143
lemma fs1_strong_cases:
+ − 144
fixes X :: "'a list"
+ − 145
shows "(X = []) \<or> (\<exists>a. \<exists> Y. (~(a memb Y) \<and> (X \<approx> a # Y)))"
+ − 146
apply (induct X)
+ − 147
apply (simp)
+ − 148
apply (metis QUOT_TYPE_I_fset.thm11 list_eq_refl mem_cons m1)
+ − 149
done
+ − 150
296
+ − 151
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
+ − 152
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
+ − 153
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
+ − 154
"IN \<equiv> membship"
163
+ − 155
+ − 156
term membship
+ − 157
term IN
+ − 158
thm IN_def
+ − 159
274
+ − 160
term fold1
+ − 161
quotient_def
+ − 162
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
+ − 163
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
+ − 164
"FOLD \<equiv> fold1"
194
+ − 165
+ − 166
term fold1
+ − 167
term fold
+ − 168
thm fold_def
+ − 169
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
+ − 170
quotient_def
254
+ − 171
fmap::"('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset"
225
+ − 172
where
254
+ − 173
"fmap \<equiv> map"
194
+ − 174
+ − 175
term map
+ − 176
term fmap
+ − 177
thm fmap_def
+ − 178
290
a0be84b0c707
removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 179
ML {* prop_of @{thm fmap_def} *}
a0be84b0c707
removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 180
274
+ − 181
ML {* val defs = @{thms EMPTY_def IN_def FUNION_def CARD_def INSERT_def fmap_def FOLD_def} *}
163
+ − 182
164
+ − 183
lemma memb_rsp:
163
+ − 184
fixes z
+ − 185
assumes a: "list_eq x y"
+ − 186
shows "(z memb x) = (z memb y)"
+ − 187
using a by induct auto
+ − 188
164
+ − 189
lemma ho_memb_rsp:
+ − 190
"(op = ===> (op \<approx> ===> op =)) (op memb) (op memb)"
214
+ − 191
by (simp add: memb_rsp)
164
+ − 192
163
+ − 193
lemma card1_rsp:
+ − 194
fixes a b :: "'a list"
+ − 195
assumes e: "a \<approx> b"
+ − 196
shows "card1 a = card1 b"
214
+ − 197
using e by induct (simp_all add:memb_rsp)
163
+ − 198
228
+ − 199
lemma ho_card1_rsp: "(op \<approx> ===> op =) card1 card1"
214
+ − 200
by (simp add: card1_rsp)
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 201
164
+ − 202
lemma cons_rsp:
163
+ − 203
fixes z
+ − 204
assumes a: "xs \<approx> ys"
+ − 205
shows "(z # xs) \<approx> (z # ys)"
+ − 206
using a by (rule list_eq.intros(5))
+ − 207
164
+ − 208
lemma ho_cons_rsp:
228
+ − 209
"(op = ===> op \<approx> ===> op \<approx>) op # op #"
214
+ − 210
by (simp add: cons_rsp)
164
+ − 211
175
+ − 212
lemma append_rsp_fst:
163
+ − 213
assumes a : "list_eq l1 l2"
214
+ − 214
shows "(l1 @ s) \<approx> (l2 @ s)"
163
+ − 215
using a
214
+ − 216
by (induct) (auto intro: list_eq.intros list_eq_refl)
+ − 217
+ − 218
lemma append_end:
+ − 219
shows "(e # l) \<approx> (l @ [e])"
+ − 220
apply (induct l)
+ − 221
apply (auto intro: list_eq.intros list_eq_refl)
+ − 222
done
+ − 223
+ − 224
lemma rev_rsp:
+ − 225
shows "a \<approx> rev a"
+ − 226
apply (induct a)
+ − 227
apply simp
+ − 228
apply (rule list_eq_refl)
+ − 229
apply simp_all
+ − 230
apply (rule list_eq.intros(6))
+ − 231
prefer 2
+ − 232
apply (rule append_rsp_fst)
+ − 233
apply assumption
+ − 234
apply (rule append_end)
+ − 235
done
163
+ − 236
214
+ − 237
lemma append_sym_rsp:
+ − 238
shows "(a @ b) \<approx> (b @ a)"
+ − 239
apply (rule list_eq.intros(6))
+ − 240
apply (rule append_rsp_fst)
+ − 241
apply (rule rev_rsp)
+ − 242
apply (rule list_eq.intros(6))
+ − 243
apply (rule rev_rsp)
+ − 244
apply (simp)
+ − 245
apply (rule append_rsp_fst)
+ − 246
apply (rule list_eq.intros(3))
+ − 247
apply (rule rev_rsp)
+ − 248
done
+ − 249
+ − 250
lemma append_rsp:
+ − 251
assumes a : "list_eq l1 r1"
+ − 252
assumes b : "list_eq l2 r2 "
+ − 253
shows "(l1 @ l2) \<approx> (r1 @ r2)"
+ − 254
apply (rule list_eq.intros(6))
+ − 255
apply (rule append_rsp_fst)
+ − 256
using a apply (assumption)
+ − 257
apply (rule list_eq.intros(6))
+ − 258
apply (rule append_sym_rsp)
+ − 259
apply (rule list_eq.intros(6))
+ − 260
apply (rule append_rsp_fst)
+ − 261
using b apply (assumption)
+ − 262
apply (rule append_sym_rsp)
+ − 263
done
175
+ − 264
194
+ − 265
lemma ho_append_rsp:
228
+ − 266
"(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"
214
+ − 267
by (simp add: append_rsp)
175
+ − 268
194
+ − 269
lemma map_rsp:
+ − 270
assumes a: "a \<approx> b"
+ − 271
shows "map f a \<approx> map f b"
+ − 272
using a
+ − 273
apply (induct)
+ − 274
apply(auto intro: list_eq.intros)
+ − 275
done
+ − 276
+ − 277
lemma ho_map_rsp:
294
+ − 278
"(op = ===> op \<approx> ===> op \<approx>) map map"
+ − 279
by (simp add: map_rsp)
194
+ − 280
294
+ − 281
lemma map_append:
258
+ − 282
"(map f (a @ b)) \<approx>
+ − 283
(map f a) @ (map f b)"
215
+ − 284
by simp (rule list_eq_refl)
194
+ − 285
273
+ − 286
lemma ho_fold_rsp:
294
+ − 287
"(op = ===> op = ===> op = ===> op \<approx> ===> op =) fold1 fold1"
292
+ − 288
apply (auto simp add: FUN_REL_EQ)
294
+ − 289
apply (case_tac "rsp_fold x")
+ − 290
prefer 2
+ − 291
apply (erule_tac list_eq.induct)
+ − 292
apply (simp_all)
+ − 293
apply (erule_tac list_eq.induct)
+ − 294
apply (simp_all)
+ − 295
apply (auto simp add: memb_rsp rsp_fold_def)
+ − 296
done
241
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 297
254
+ − 298
print_quotients
+ − 299
+ − 300
226
+ − 301
ML {* val qty = @{typ "'a fset"} *}
+ − 302
ML {* val rsp_thms =
273
+ − 303
@{thms ho_memb_rsp ho_cons_rsp ho_card1_rsp ho_map_rsp ho_append_rsp ho_fold_rsp}
226
+ − 304
@ @{thms ho_all_prs ho_ex_prs} *}
206
+ − 305
239
+ − 306
ML {* fun lift_thm_fset lthy t = lift_thm lthy qty "fset" rsp_thms defs t *}
348
+ − 307
ML {* fun lift_thm_g_fset lthy t g = lift_thm_goal lthy qty "fset" rsp_thms defs t g *}
364
+ − 308
ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
+ − 309
ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "fset"; *}
+ − 310
ML {* val consts = lookup_quot_consts defs *}
+ − 311
ML {* fun lift_tac_fset lthy t = lift_tac lthy t rel_eqv rel_refl rty quot trans2 rsp_thms reps_same defs *}
314
+ − 312
364
+ − 313
lemma "IN x EMPTY = False"
+ − 314
by (tactic {* lift_tac_fset @{context} @{thm m1} 1 *})
353
9a0e8ab42ee8
fixed the error by a temporary fix (the data of the eqivalence relation should be only its name)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 315
364
+ − 316
lemma "IN x (INSERT y xa) = (x = y \<or> IN x xa)"
+ − 317
by (tactic {* lift_tac_fset @{context} @{thm m2} 1 *})
356
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 318
364
+ − 319
lemma "INSERT a (INSERT a x) = INSERT a x"
+ − 320
apply (tactic {* lift_tac_fset @{context} @{thm list_eq.intros(4)} 1 *})
+ − 321
done
+ − 322
+ − 323
lemma "x = xa \<longrightarrow> INSERT a x = INSERT a xa"
+ − 324
apply (tactic {* lift_tac_fset @{context} @{thm list_eq.intros(5)} 1 *})
+ − 325
done
353
9a0e8ab42ee8
fixed the error by a temporary fix (the data of the eqivalence relation should be only its name)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 326
364
+ − 327
lemma "CARD x = Suc n \<longrightarrow> (\<exists>a b. \<not> IN a b & x = INSERT a b)"
+ − 328
apply (tactic {* lift_tac_fset @{context} @{thm card1_suc} 1 *})
+ − 329
done
+ − 330
+ − 331
lemma "(\<not> IN x xa) = (CARD (INSERT x xa) = Suc (CARD xa))"
+ − 332
apply (tactic {* lift_tac_fset @{context} @{thm not_mem_card1} 1 *})
+ − 333
done
356
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 334
364
+ − 335
ML {* atomize_thm @{thm fold1.simps(2)} *}
+ − 336
(*ML {* lift_thm_g_fset @{context} @{thm fold1.simps(2)} @{term "FOLD f g (z::'b) (INSERT a x) =
+ − 337
(if rsp_fold f then if IN a x then FOLD f g z x else f (g a) (FOLD f g z x) else z)"} *}*)
356
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 338
364
+ − 339
lemma "\<forall>f g z a x. FOLD f g (z::'b) (INSERT a x) =
+ − 340
(if rsp_fold f then if IN a x then FOLD f g z x else f (g a) (FOLD f g z x) else z)"
+ − 341
apply (tactic {* lift_tac_fset @{context} @{thm fold1.simps(2)} 1 *})
+ − 342
done
356
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 343
364
+ − 344
ML {* cterm_of @{theory} (atomize_goal @{theory} @{term "\<forall>x xa a. x = xa \<Longrightarrow> INSERT a x = INSERT a xa"}) *}
356
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 345
348
+ − 346
(* Doesn't work with 'a, 'b, but works with 'b, 'a *)
364
+ − 347
ML {* lift_thm_g_fset @{context} @{thm fold1.simps(2)}
348
+ − 348
+ − 349
ML {* lift_thm_fset @{context} @{thm append_assoc} *}
356
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 350
ML {* lift_thm_g_fset @{context} @{thm append_assoc} @{term "FUNION (FUNION x xa) xb = FUNION x (FUNION xa xb)"} *}
353
9a0e8ab42ee8
fixed the error by a temporary fix (the data of the eqivalence relation should be only its name)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 351
356
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 352
ML {* lift_thm_fset @{context} @{thm map_append} *}
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 353
ML {* lift_thm_g_fset @{context} @{thm map_append} @{term "fmap f (FUNION (x::'b fset) (xa::'b fset)) = FUNION (fmap f x) (fmap f xa)"} *}
349
+ − 354
ML {* lift_thm_fset @{context} @{thm list.induct} *}
+ − 355
ML {* lift_thm_g_fset @{context} @{thm list.induct} @{term "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l"} *}
348
+ − 356
364
+ − 357
+ − 358
+ − 359
apply(tactic {* procedure_tac @{thm m1} @{context} 1 *})
+ − 360
apply(tactic {* regularize_tac @{context} rel_eqv rel_refl 1 *})
+ − 361
apply(tactic {* REPEAT_ALL_NEW (r_mk_comb_tac @{context} rty quot rel_refl trans2 rsp_thms) 1 *})
+ − 362
apply(tactic {* clean_tac @{context} quot defs reps_same 1 *})
+ − 363
apply(tactic {* TRY' (REPEAT_ALL_NEW (allex_prs_tac @{context} quot)) 1 *})
+ − 364
apply(tactic {* TRY' (REPEAT_ALL_NEW (lambda_prs_tac @{context} quot)) 1 *})
+ − 365
ML_prf {* val lower = flat (map (add_lower_defs @{context}) defs) *}
+ − 366
apply(tactic {* REPEAT_ALL_NEW (EqSubst.eqsubst_tac @{context} [0] lower) 1*})
+ − 367
apply(tactic {* simp_tac (HOL_ss addsimps [reps_same]) 1 *})
+ − 368
+ − 369
+ − 370
+ − 371
348
+ − 372
(*ML {* lift_thm_fset @{context} @{thm neq_Nil_conv} *}*)
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 373
273
+ − 374
quotient_def
276
+ − 375
fset_rec::"'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
273
+ − 376
where
+ − 377
"fset_rec \<equiv> list_rec"
+ − 378
292
+ − 379
quotient_def
+ − 380
fset_case::"'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
+ − 381
where
+ − 382
"fset_case \<equiv> list_case"
+ − 383
296
+ − 384
(* Probably not true without additional assumptions about the function *)
292
+ − 385
lemma list_rec_rsp:
+ − 386
"(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_rec list_rec"
+ − 387
apply (auto simp add: FUN_REL_EQ)
296
+ − 388
apply (erule_tac list_eq.induct)
+ − 389
apply (simp_all)
292
+ − 390
sorry
289
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 391
292
+ − 392
lemma list_case_rsp:
+ − 393
"(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_case list_case"
+ − 394
apply (auto simp add: FUN_REL_EQ)
+ − 395
sorry
+ − 396
+ − 397
+ − 398
ML {* val rsp_thms = @{thms list_rec_rsp list_case_rsp} @ rsp_thms *}
+ − 399
ML {* val defs = @{thms fset_rec_def fset_case_def} @ defs *}
+ − 400
+ − 401
ML {* fun lift_thm_fset lthy t = lift_thm lthy qty "fset" rsp_thms defs t *}
356
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 402
ML {* fun lift_thm_g_fset lthy t g = lift_thm_goal lthy qty "fset" rsp_thms defs t g *}
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 403
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 404
thm list.recs(2)
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 405
ML {* lift_thm_fset @{context} @{thm list.recs(2)} *}
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 406
ML {* lift_thm_g_fset @{context} @{thm list.recs(2)} @{term "fset_rec (f1::'t) x (INSERT a xa) = x a xa (fset_rec f1 x xa)"} *}
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 407
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 408
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 409
ML {* val gl = @{term "fset_rec (f1::'t) x (INSERT a xa) = x a xa (fset_rec f1 x xa)"} *}
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 410
ML {* val t_a = atomize_thm @{thm list.recs(2)} *}
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 411
ML {* val qtrm = atomize_goal @{theory} gl *}
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 412
ML {* val rg = cterm_of @{theory}(Syntax.check_term @{context} (REGULARIZE_trm @{context} (prop_of t_a) qtrm)) *}
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 413
ML {* val rg2 = cterm_of @{theory}(my_reg @{context} rel rty (prop_of t_a)) *}
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 414
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 415
ML {* val t_r = regularize_goal @{context} t_a rel_eqv rel_refl qtrm *}
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 416
ML {* val t_r = regularize t_a rty rel rel_eqv rel_refl @{context} *}
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 417
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 418
ML {* val rg = cterm_of @{theory}(Syntax.check_term @{context} (inj_REPABS @{context} ((prop_of t_r), qtrm))) *}
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 419
ML {* val rg2 = cterm_of @{theory} (build_repabs_term @{context} t_r consts rty qty) *}
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 420
ML {* val t_t = repabs_goal @{context} t_r rty quot rel_refl trans2 rsp_thms qtrm *}
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 421
ML {* val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms *}
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 422
ML {*
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 423
val lthy = @{context}
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 424
val (alls, exs) = findallex lthy rty qty (prop_of t_a);
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 425
val allthms = map (make_allex_prs_thm lthy quot @{thm FORALL_PRS}) alls
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 426
val exthms = map (make_allex_prs_thm lthy quot @{thm EXISTS_PRS}) exs
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 427
val t_a = MetaSimplifier.rewrite_rule (allthms @ exthms) t_t
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 428
val abs = findabs rty (prop_of t_a);
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 429
val aps = findaps rty (prop_of t_a);
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 430
val app_prs_thms = map (applic_prs lthy rty qty absrep) aps;
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 431
val lam_prs_thms = map (make_simp_prs_thm lthy quot @{thm LAMBDA_PRS}) abs;
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 432
val t_l = repeat_eqsubst_thm lthy (lam_prs_thms @ app_prs_thms) t_a;
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 433
val defs_sym = flat (map (add_lower_defs lthy) defs);
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 434
val defs_sym_eq = map (fn x => eq_reflection OF [x]) defs_sym;
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 435
val t_id = simp_ids lthy t_l;
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 436
val t_d0 = MetaSimplifier.rewrite_rule defs_sym_eq t_id;
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 437
val t_d = repeat_eqsubst_thm lthy defs_sym t_d0;
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 438
val t_r = MetaSimplifier.rewrite_rule [reps_same] t_d;
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 439
val t_rv = ObjectLogic.rulify t_r
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 440
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 441
*}
292
+ − 442
300
+ − 443
356
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 444
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 445
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 446
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 447
292
+ − 448
ML {* map (lift_thm_fset @{context}) @{thms list.cases} *}
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 449
356
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 450
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 451
51aafebf4d06
Another theorem for which the new regularize differs from old one, so the goal is not proved. But it seems, that the new one is better.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 452
348
+ − 453
ML {* atomize_thm @{thm m1} *}
+ − 454
ML {* cterm_of @{theory} (atomize_goal @{theory} @{term "IN x EMPTY = False"}) *}
+ − 455
ML {* lift_thm_fset @{context} @{thm m1} *}
+ − 456
(* ML {* lift_thm_g_fset @{context} @{thm m1} @{term "IN x EMPTY = False"}) *} *)
+ − 457
+ − 458
304
+ − 459
lemma list_induct_part:
+ − 460
assumes a: "P (x :: 'a list) ([] :: 'a list)"
+ − 461
assumes b: "\<And>e t. P x t \<Longrightarrow> P x (e # t)"
+ − 462
shows "P x l"
+ − 463
apply (rule_tac P="P x" in list.induct)
+ − 464
apply (rule a)
+ − 465
apply (rule b)
+ − 466
apply (assumption)
+ − 467
done
273
+ − 468
292
+ − 469
+ − 470
(* Construction site starts here *)
+ − 471
+ − 472
273
+ − 473
ML {* val consts = lookup_quot_consts defs *}
+ − 474
ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
+ − 475
ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "fset" *}
+ − 476
296
+ − 477
thm list.recs(2)
304
+ − 478
ML {* val t_a = atomize_thm @{thm list_induct_part} *}
334
+ − 479
+ − 480
285
+ − 481
(* prove {* build_regularize_goal t_a rty rel @{context} *}
+ − 482
ML_prf {* fun tac ctxt = FIRST' [
251
+ − 483
rtac rel_refl,
+ − 484
atac,
285
+ − 485
rtac @{thm universal_twice},
+ − 486
(rtac @{thm impI} THEN' atac),
+ − 487
rtac @{thm implication_twice},
334
+ − 488
//comented out rtac @{thm equality_twice}, //
285
+ − 489
EqSubst.eqsubst_tac ctxt [0]
239
+ − 490
[(@{thm equiv_res_forall} OF [rel_eqv]),
285
+ − 491
(@{thm equiv_res_exists} OF [rel_eqv])],
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 492
(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
+ − 493
(rtac @{thm RIGHT_RES_FORALL_REGULAR})
285
+ − 494
]; *}
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 495
apply (atomize(full))
285
+ − 496
apply (tactic {* REPEAT_ALL_NEW (tac @{context}) 1 *})
334
+ − 497
done *)
305
+ − 498
ML {* val t_r = regularize t_a rty rel rel_eqv rel_refl @{context} *}
273
+ − 499
ML {*
285
+ − 500
val rt = build_repabs_term @{context} t_r consts rty qty
+ − 501
val rg = Logic.mk_equals ((Thm.prop_of t_r), rt);
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 502
*}
300
+ − 503
prove {* Syntax.check_term @{context} rg *}
309
+ − 504
ML_prf {* fun r_mk_comb_tac_fset lthy = r_mk_comb_tac lthy rty 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
+ − 505
apply(atomize(full))
194
+ − 506
apply (tactic {* REPEAT_ALL_NEW (r_mk_comb_tac_fset @{context}) 1 *})
273
+ − 507
done
305
+ − 508
ML {*
+ − 509
val t_t = repabs @{context} t_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
+ − 510
*}
241
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 511
292
+ − 512
ML {* val abs = findabs rty (prop_of (t_a)) *}
+ − 513
ML {* val aps = findaps rty (prop_of (t_a)) *}
285
+ − 514
ML {* val lam_prs_thms = map (make_simp_prs_thm @{context} quot @{thm LAMBDA_PRS}) abs *}
+ − 515
ML {* val app_prs_thms = map (applic_prs @{context} rty qty absrep) aps *}
289
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 516
ML {* val lam_prs_thms = map Thm.varifyT lam_prs_thms *}
292
+ − 517
ML {* t_t *}
289
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 518
ML {* val (alls, exs) = findallex @{context} rty qty (prop_of t_a); *}
285
+ − 519
ML {* val allthms = map (make_allex_prs_thm @{context} quot @{thm FORALL_PRS}) alls *}
304
+ − 520
ML {* val t_l0 = repeat_eqsubst_thm @{context} (app_prs_thms) t_t *}
+ − 521
ML app_prs_thms
296
+ − 522
ML {* val t_l = repeat_eqsubst_thm @{context} (lam_prs_thms) t_l0 *}
304
+ − 523
ML lam_prs_thms
+ − 524
ML {* val t_id = simp_ids @{context} t_l *}
+ − 525
thm INSERT_def
292
+ − 526
ML {* val defs_sym = flat (map (add_lower_defs @{context}) defs) *}
+ − 527
ML {* val t_d = repeat_eqsubst_thm @{context} defs_sym t_id *}
304
+ − 528
ML allthms
+ − 529
thm FORALL_PRS
309
+ − 530
ML {* val t_al = MetaSimplifier.rewrite_rule (allthms) t_d *}
+ − 531
ML {* val t_s = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.REPS_same} t_al *}
304
+ − 532
ML {* ObjectLogic.rulify t_s *}
+ − 533
+ − 534
ML {* val gl = @{term "P (x :: 'a list) (EMPTY :: 'a fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"} *}
338
+ − 535
ML {* val gla = atomize_goal @{theory} gl *}
309
+ − 536
338
+ − 537
prove t_r: {* mk_REGULARIZE_goal @{context} (prop_of t_a) gla *}
332
+ − 538
309
+ − 539
ML_prf {* fun tac ctxt = FIRST' [
+ − 540
rtac rel_refl,
+ − 541
atac,
+ − 542
rtac @{thm universal_twice},
+ − 543
(rtac @{thm impI} THEN' atac),
+ − 544
rtac @{thm implication_twice},
+ − 545
(*rtac @{thm equality_twice},*)
+ − 546
EqSubst.eqsubst_tac ctxt [0]
+ − 547
[(@{thm equiv_res_forall} OF [rel_eqv]),
+ − 548
(@{thm equiv_res_exists} OF [rel_eqv])],
+ − 549
(rtac @{thm impI} THEN' (asm_full_simp_tac (Simplifier.context ctxt HOL_ss)) THEN' rtac rel_refl),
+ − 550
(rtac @{thm RIGHT_RES_FORALL_REGULAR})
+ − 551
]; *}
+ − 552
+ − 553
apply (atomize(full))
+ − 554
apply (tactic {* REPEAT_ALL_NEW (tac @{context}) 1 *})
+ − 555
done
+ − 556
+ − 557
ML {* val t_r = @{thm t_r} OF [t_a] *}
+ − 558
338
+ − 559
ML {* val ttt = mk_inj_REPABS_goal @{context} (prop_of t_r, gla) *}
309
+ − 560
ML {* val si = simp_ids_trm (cterm_of @{theory} ttt) *}
337
553bef083318
Removed second implementation of Regularize/Inject from FSet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 561
prove t_t: {* term_of si *}
309
+ − 562
ML_prf {* fun r_mk_comb_tac_fset lthy = r_mk_comb_tac lthy rty quot rel_refl trans2 rsp_thms *}
+ − 563
apply(atomize(full))
+ − 564
apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *})
+ − 565
apply (rule FUN_QUOTIENT)
+ − 566
apply (rule FUN_QUOTIENT)
+ − 567
apply (rule IDENTITY_QUOTIENT)
+ − 568
apply (rule FUN_QUOTIENT)
+ − 569
apply (rule QUOTIENT_fset)
+ − 570
apply (rule IDENTITY_QUOTIENT)
+ − 571
apply (rule IDENTITY_QUOTIENT)
+ − 572
apply (rule IDENTITY_QUOTIENT)
+ − 573
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 574
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 575
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 576
apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *})
+ − 577
apply (rule IDENTITY_QUOTIENT)
+ − 578
apply (rule IDENTITY_QUOTIENT)
+ − 579
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 580
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 581
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 582
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 583
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 584
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 585
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 586
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 587
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 588
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 589
apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *})
+ − 590
apply (rule IDENTITY_QUOTIENT)
+ − 591
apply (rule FUN_QUOTIENT)
+ − 592
apply (rule QUOTIENT_fset)
+ − 593
apply (rule IDENTITY_QUOTIENT)
+ − 594
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 595
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 596
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 597
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 598
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 599
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 600
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 601
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 602
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 603
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 604
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 605
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 606
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 607
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 608
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 609
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 610
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 611
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
317
+ − 612
apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
+ − 613
apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
+ − 614
apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP(1)} @{context} THEN' (RANGE [quotient_tac quot])) 1 *})
+ − 615
apply assumption
+ − 616
apply (rule refl)
309
+ − 617
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 618
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
317
+ − 619
apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
+ − 620
apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
+ − 621
apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP(1)} @{context} THEN' (RANGE [quotient_tac quot])) 1 *})
309
+ − 622
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 623
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 624
apply (tactic {* REPEAT_ALL_NEW (r_mk_comb_tac_fset @{context}) 1 *})
+ − 625
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
317
+ − 626
apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
+ − 627
apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP(1)} @{context} THEN' (RANGE [quotient_tac quot])) 1 *})
309
+ − 628
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 629
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 630
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 631
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
+ − 632
done
+ − 633
+ − 634
thm t_t
+ − 635
ML {* val t_t = @{thm Pure.equal_elim_rule1} OF [@{thm t_t}, t_r] *}
+ − 636
ML {* val t_l = repeat_eqsubst_thm @{context} (lam_prs_thms) t_t *}
+ − 637
ML {* val t_d = repeat_eqsubst_thm @{context} defs_sym t_l *}
+ − 638
ML {* val t_al = MetaSimplifier.rewrite_rule (allthms) t_d *}
+ − 639
ML {* val t_s = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.REPS_same} t_al *}
+ − 640
178
945786a68ec6
Finally lifted induction, with some manually added simplification lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 641
163
+ − 642
ML {*
226
+ − 643
fun lift_thm_fset_note name thm lthy =
163
+ − 644
let
226
+ − 645
val lifted_thm = lift_thm_fset lthy thm;
163
+ − 646
val (_, lthy2) = note (name, lifted_thm) lthy;
+ − 647
in
+ − 648
lthy2
+ − 649
end;
+ − 650
*}
+ − 651
226
+ − 652
local_setup {*
+ − 653
lift_thm_fset_note @{binding "m1l"} @{thm m1} #>
+ − 654
lift_thm_fset_note @{binding "m2l"} @{thm m2} #>
+ − 655
lift_thm_fset_note @{binding "leqi4l"} @{thm list_eq.intros(4)} #>
+ − 656
lift_thm_fset_note @{binding "leqi5l"} @{thm list_eq.intros(5)}
163
+ − 657
*}
226
+ − 658
thm m1l
+ − 659
thm m2l
+ − 660
thm leqi4l
+ − 661
thm leqi5l
163
+ − 662
+ − 663
end