src/Pure/Isar/local_defs.ML
author wenzelm
Wed, 24 Jun 2009 21:28:02 +0200
changeset 31794 71af1fd6a5e4
parent 30770 6976521b4263
child 32111 30e2ffbba718
permissions -rw-r--r--
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
renamed Variable.importT_thms to Variable.importT (again);
     1 (*  Title:      Pure/Isar/local_defs.ML
     2     Author:     Makarius
     3 
     4 Local definitions.
     5 *)
     6 
     7 signature LOCAL_DEFS =
     8 sig
     9   val cert_def: Proof.context -> term -> (string * typ) * term
    10   val abs_def: term -> (string * typ) * term
    11   val mk_def: Proof.context -> (string * term) list -> term list
    12   val expand: cterm list -> thm -> thm
    13   val def_export: Assumption.export
    14   val add_defs: ((binding * mixfix) * (Thm.binding * term)) list -> Proof.context ->
    15     (term * (string * thm)) list * Proof.context
    16   val add_def: (binding * mixfix) * term -> Proof.context -> (term * thm) * Proof.context
    17   val fixed_abbrev: (binding * mixfix) * term -> Proof.context ->
    18     (term * term) * Proof.context
    19   val export: Proof.context -> Proof.context -> thm -> thm list * thm
    20   val export_cterm: Proof.context -> Proof.context -> cterm -> cterm * thm
    21   val trans_terms: Proof.context -> thm list -> thm
    22   val trans_props: Proof.context -> thm list -> thm
    23   val print_rules: Proof.context -> unit
    24   val defn_add: attribute
    25   val defn_del: attribute
    26   val meta_rewrite_conv: Proof.context -> conv
    27   val meta_rewrite_rule: Proof.context -> thm -> thm
    28   val unfold: Proof.context -> thm list -> thm -> thm
    29   val unfold_goals: Proof.context -> thm list -> thm -> thm
    30   val unfold_tac: Proof.context -> thm list -> tactic
    31   val fold: Proof.context -> thm list -> thm -> thm
    32   val fold_tac: Proof.context -> thm list -> tactic
    33   val derived_def: Proof.context -> bool -> term ->
    34     ((string * typ) * term) * (Proof.context -> thm -> thm)
    35 end;
    36 
    37 structure LocalDefs: LOCAL_DEFS =
    38 struct
    39 
    40 (** primitive definitions **)
    41 
    42 (* prepare defs *)
    43 
    44 fun cert_def ctxt eq =
    45   let
    46     fun err msg = cat_error msg ("The error(s) above occurred in definition: " ^
    47       quote (Syntax.string_of_term ctxt eq));
    48     val ((lhs, _), eq') = eq
    49       |> Sign.no_vars (Syntax.pp ctxt)
    50       |> PrimitiveDefs.dest_def ctxt Term.is_Free (Variable.is_fixed ctxt) (K true)
    51       handle TERM (msg, _) => err msg | ERROR msg => err msg;
    52   in (Term.dest_Free (Term.head_of lhs), eq') end;
    53 
    54 val abs_def = PrimitiveDefs.abs_def #>> Term.dest_Free;
    55 
    56 fun mk_def ctxt args =
    57   let
    58     val (xs, rhss) = split_list args;
    59     val (bind, _) = ProofContext.bind_fixes xs ctxt;
    60     val lhss = map (fn (x, rhs) => bind (Free (x, Term.fastype_of rhs))) args;
    61   in map Logic.mk_equals (lhss ~~ rhss) end;
    62 
    63 
    64 (* export defs *)
    65 
    66 val head_of_def =
    67   #1 o Term.dest_Free o Term.head_of o #1 o Logic.dest_equals o Term.strip_all_body;
    68 
    69 
    70 (*
    71   [x, x == a]
    72        :
    73       B x
    74   -----------
    75       B a
    76 *)
    77 fun expand defs =
    78   Drule.implies_intr_list defs
    79   #> Drule.generalize ([], map (head_of_def o Thm.term_of) defs)
    80   #> funpow (length defs) (fn th => Drule.reflexive_thm RS th);
    81 
    82 val expand_term = Envir.expand_term_frees o map (abs_def o Thm.term_of);
    83 
    84 fun def_export _ defs = (expand defs, expand_term defs);
    85 
    86 
    87 (* add defs *)
    88 
    89 fun add_defs defs ctxt =
    90   let
    91     val ((bvars, mxs), specs) = defs |> split_list |>> split_list;
    92     val ((bfacts, atts), rhss) = specs |> split_list |>> split_list;
    93     val xs = map Name.of_binding bvars;
    94     val names = map2 Thm.def_binding_optional bvars bfacts;
    95     val eqs = mk_def ctxt (xs ~~ rhss);
    96     val lhss = map (fst o Logic.dest_equals) eqs;
    97   in
    98     ctxt
    99     |> ProofContext.add_fixes (map2 (fn x => fn mx => (x, NONE, mx)) bvars mxs) |> #2
   100     |> fold Variable.declare_term eqs
   101     |> ProofContext.add_assms_i def_export
   102       (map2 (fn a => fn eq => (a, [(eq, [])])) (names ~~ atts) eqs)
   103     |>> map2 (fn lhs => fn (name, [th]) => (lhs, (name, th))) lhss
   104   end;
   105 
   106 fun add_def (var, rhs) ctxt =
   107   let val ([(lhs, (_, th))], ctxt') = add_defs [(var, (Thm.empty_binding, rhs))] ctxt
   108   in ((lhs, th), ctxt') end;
   109 
   110 
   111 (* fixed_abbrev *)
   112 
   113 fun fixed_abbrev ((x, mx), rhs) ctxt =
   114   let
   115     val T = Term.fastype_of rhs;
   116     val ([x'], ctxt') = ctxt
   117       |> Variable.declare_term rhs
   118       |> ProofContext.add_fixes [(x, SOME T, mx)];
   119     val lhs = Free (x', T);
   120     val _ = cert_def ctxt' (Logic.mk_equals (lhs, rhs));
   121     fun abbrev_export _ _ = (I, Envir.expand_term_frees [((x', T), rhs)]);
   122     val (_, ctxt'') = Assumption.add_assms abbrev_export [] ctxt';
   123   in ((lhs, rhs), ctxt'') end;
   124 
   125 
   126 (* specific export -- result based on educated guessing *)
   127 
   128 (*
   129   [xs, xs == as]
   130         :
   131        B xs
   132   --------------
   133        B as
   134 *)
   135 fun export inner outer =    (*beware of closure sizes*)
   136   let
   137     val exp = Assumption.export false inner outer;
   138     val prems = Assumption.all_prems_of inner;
   139   in fn th =>
   140     let
   141       val th' = exp th;
   142       val still_fixed = map #1 (Thm.fold_terms Term.add_frees th' []);
   143       val defs = prems |> filter_out (fn prem =>
   144         (case try (head_of_def o Thm.prop_of) prem of
   145           SOME x => member (op =) still_fixed x
   146         | NONE => true));
   147     in (map Drule.abs_def defs, th') end
   148   end;
   149 
   150 (*
   151   [xs, xs == as]
   152         :
   153      TERM b xs
   154   --------------  and  --------------
   155      TERM b as          b xs == b as
   156 *)
   157 fun export_cterm inner outer ct =
   158   let val (defs, ct') = export inner outer (Drule.mk_term ct) ||> Drule.dest_term
   159   in (ct', MetaSimplifier.rewrite true defs ct) end;
   160 
   161 
   162 (* basic transitivity reasoning -- modulo beta-eta *)
   163 
   164 local
   165 
   166 val is_trivial = Pattern.aeconv o Logic.dest_equals o Thm.prop_of;
   167 
   168 fun trans_list _ _ [] = raise Empty
   169   | trans_list trans ctxt (th :: raw_eqs) =
   170       (case filter_out is_trivial raw_eqs of
   171         [] => th
   172       | eqs =>
   173           let val ((_, th' :: eqs'), ctxt') = Variable.import true (th :: eqs) ctxt
   174           in singleton (Variable.export ctxt' ctxt) (fold trans eqs' th') end);
   175 
   176 in
   177 
   178 val trans_terms = trans_list (fn eq2 => fn eq1 => eq2 COMP (eq1 COMP Drule.transitive_thm));
   179 val trans_props = trans_list (fn eq => fn th => th COMP (eq COMP Drule.equal_elim_rule1));
   180 
   181 end;
   182 
   183 
   184 
   185 (** defived definitions **)
   186 
   187 (* transformation rules *)
   188 
   189 structure Rules = GenericDataFun
   190 (
   191   type T = thm list;
   192   val empty = []
   193   val extend = I;
   194   fun merge _ = Thm.merge_thms;
   195 );
   196 
   197 fun print_rules ctxt =
   198   Pretty.writeln (Pretty.big_list "definitional transformations:"
   199     (map (ProofContext.pretty_thm ctxt) (Rules.get (Context.Proof ctxt))));
   200 
   201 val defn_add = Thm.declaration_attribute (Rules.map o Thm.add_thm);
   202 val defn_del = Thm.declaration_attribute (Rules.map o Thm.del_thm);
   203 
   204 
   205 (* meta rewrite rules *)
   206 
   207 fun meta_rewrite_conv ctxt =
   208   MetaSimplifier.rewrite_cterm (false, false, false) (K (K NONE))
   209     (MetaSimplifier.context ctxt MetaSimplifier.empty_ss
   210       addeqcongs [Drule.equals_cong]    (*protect meta-level equality*)
   211       addsimps (Rules.get (Context.Proof ctxt)));
   212 
   213 val meta_rewrite_rule = Conv.fconv_rule o meta_rewrite_conv;
   214 
   215 
   216 (* rewriting with object-level rules *)
   217 
   218 fun meta f ctxt = f o map (meta_rewrite_rule ctxt);
   219 
   220 val unfold       = meta MetaSimplifier.rewrite_rule;
   221 val unfold_goals = meta MetaSimplifier.rewrite_goals_rule;
   222 val unfold_tac   = meta MetaSimplifier.rewrite_goals_tac;
   223 val fold         = meta MetaSimplifier.fold_rule;
   224 val fold_tac     = meta MetaSimplifier.fold_goals_tac;
   225 
   226 
   227 (* derived defs -- potentially within the object-logic *)
   228 
   229 fun derived_def ctxt conditional prop =
   230   let
   231     val ((c, T), rhs) = prop
   232       |> Thm.cterm_of (ProofContext.theory_of ctxt)
   233       |> meta_rewrite_conv ctxt
   234       |> (snd o Logic.dest_equals o Thm.prop_of)
   235       |> conditional ? Logic.strip_imp_concl
   236       |> (abs_def o #2 o cert_def ctxt);
   237     fun prove ctxt' def =
   238       let
   239         val frees = Term.fold_aterms (fn Free (x, _) =>
   240           if Variable.is_fixed ctxt' x then I else insert (op =) x | _ => I) prop [];
   241       in
   242         Goal.prove ctxt' frees [] prop (K (ALLGOALS
   243           (CONVERSION (meta_rewrite_conv ctxt') THEN'
   244             MetaSimplifier.rewrite_goal_tac [def] THEN'
   245             resolve_tac [Drule.reflexive_thm])))
   246         handle ERROR msg => cat_error msg "Failed to prove definitional specification."
   247       end;
   248   in (((c, T), rhs), prove) end;
   249 
   250 end;