src/Pure/Isar/local_defs.ML
changeset 7502 257dd7777676
parent 7416 a98963d70f81
child 7667 22dc8b2455b8
equal deleted inserted replaced
7501:76ed51454609 7502:257dd7777676
    15 
    15 
    16 structure LocalDefs: LOCAL_DEFS =
    16 structure LocalDefs: LOCAL_DEFS =
    17 struct
    17 struct
    18 
    18 
    19 
    19 
    20 val refl_tac = Tactic.rtac (standard (Drule.reflexive_thm RS Drule.triv_goal));
    20 val refl_tac = Tactic.rtac (Drule.standard (Drule.reflexive_thm RS Drule.triv_goal));
    21 
    21 
    22 
    22 
    23 fun gen_def fix prep_term match_binds raw_name atts ((x, raw_T), (raw_rhs, raw_pats)) state =
    23 fun gen_def fix prep_term match_binds raw_name atts ((x, raw_T), (raw_rhs, raw_pats)) state =
    24   let
    24   let
    25     fun err msg = raise Proof.STATE ("Bad local def: " ^ msg, state);
    25     fun err msg = raise Proof.STATE ("Bad local def: " ^ msg, state);