| 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
 | 
|  |      9 |   val eq_reflection = eq_reflection
 | 
| 9528 |     10 |   val rev_eq_reflection = meta_eq_to_obj_eq
 | 
| 7355 |     11 |   val imp_intr = impI
 | 
|  |     12 |   val rev_mp = rev_mp
 | 
|  |     13 |   val subst = subst
 | 
|  |     14 |   val sym = sym
 | 
| 20974 |     15 |   val thin_refl = prove_goal (the_context ()) "!!X. [|x=x; PROP W|] ==> PROP W" (K [atac 1]);
 | 
|  |     16 | end;
 | 
| 7355 |     17 | 
 | 
|  |     18 | structure Hypsubst = HypsubstFun(Hypsubst_Data);
 | 
|  |     19 | open Hypsubst;
 |