author | nipkow |
Tue, 01 Sep 2009 19:48:11 +0200 | |
changeset 32475 | d2c97fc18704 |
parent 27572 | 67cd6ed76446 |
child 39159 | 0dec18004e75 |
permissions | -rw-r--r-- |
7355 | 1 |
|
2 |
(** Applying HypsubstFun to generate hyp_subst_tac **) |
|
3 |
structure Hypsubst_Data = |
|
20974 | 4 |
struct |
7355 | 5 |
structure Simplifier = Simplifier |
21218 | 6 |
val dest_eq = FOLogic.dest_eq |
7355 | 7 |
val dest_Trueprop = FOLogic.dest_Trueprop |
8 |
val dest_imp = FOLogic.dest_imp |
|
21539 | 9 |
val eq_reflection = thm "eq_reflection" |
10 |
val rev_eq_reflection = thm "meta_eq_to_obj_eq" |
|
11 |
val imp_intr = thm "impI" |
|
12 |
val rev_mp = thm "rev_mp" |
|
13 |
val subst = thm "subst" |
|
14 |
val sym = thm "sym" |
|
15 |
val thin_refl = thm "thin_refl" |
|
27572
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
krauss
parents:
21539
diff
changeset
|
16 |
val prop_subst = @{lemma "PROP P(t) ==> PROP prop (x = t ==> PROP P(x))" |
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
krauss
parents:
21539
diff
changeset
|
17 |
by (unfold prop_def) (drule eq_reflection, unfold)} |
20974 | 18 |
end; |
7355 | 19 |
|
20 |
structure Hypsubst = HypsubstFun(Hypsubst_Data); |
|
21 |
open Hypsubst; |