Nominal-General/nominal_eqvt.ML
changeset 1835 636de31888a6
parent 1833 2050b5723c04
child 1861 226b797868dc
equal deleted inserted replaced
1834:9909cc3566c5 1835:636de31888a6
     1 (*  Title:      nominal_eqvt.ML
     1 (*  Title:      nominal_eqvt.ML
     2     Author:     Stefan Berghofer
     2     Author:     Stefan Berghofer (original code)
     3     Author:     Christian Urban
     3     Author:     Christian Urban
     4 
     4 
     5     Automatic proofs for equivariance of inductive predicates.
     5     Automatic proofs for equivariance of inductive predicates.
     6 *)
     6 *)
     7 
     7 
     8 signature NOMINAL_EQVT =
     8 signature NOMINAL_EQVT =
     9 sig
     9 sig
    10   val eqvt_rel_tac : xstring -> Proof.context -> local_theory
    10   val equivariance: string -> Proof.context -> local_theory
       
    11   val eqvt_rel_tac: Proof.context -> string list -> term -> thm -> thm list -> int -> tactic
       
    12   val eqvt_rel_case_tac: Proof.context -> string list -> term -> thm -> int -> tactic
    11 end
    13 end
    12 
    14 
    13 structure Nominal_Eqvt : NOMINAL_EQVT =
    15 structure Nominal_Eqvt : NOMINAL_EQVT =
    14 struct
    16 struct
    15 
    17 
    21     (HOL_basic_ss addsimps @{thms induct_atomize});
    23     (HOL_basic_ss addsimps @{thms induct_atomize});
    22 val atomize_intr = Conv.fconv_rule (Conv.prems_conv ~1 atomize_conv);
    24 val atomize_intr = Conv.fconv_rule (Conv.prems_conv ~1 atomize_conv);
    23 fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1
    25 fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1
    24   (Conv.params_conv ~1 (K (Conv.prems_conv ~1 atomize_conv)) ctxt));
    26   (Conv.params_conv ~1 (K (Conv.prems_conv ~1 atomize_conv)) ctxt));
    25 
    27 
       
    28 
       
    29 (** 
       
    30  proves F[f t] from F[t] which is the given theorem  
       
    31   - F needs to be monotone
       
    32   - f returns either SOME for a term it fires 
       
    33     and NONE elsewhere 
       
    34 **)
    26 fun map_term f t = 
    35 fun map_term f t = 
    27   (case f t of
    36   (case f t of
    28      NONE => map_term' f t 
    37      NONE => map_term' f t 
    29    | x => x)
    38    | x => x)
    30 and map_term' f (t $ u) = 
    39 and map_term' f (t $ u) = 
    46   EVERY [cut_facts_tac [thm] 1, etac rev_mp 1,
    55   EVERY [cut_facts_tac [thm] 1, etac rev_mp 1,
    47     REPEAT_DETERM (FIRSTGOAL (resolve_tac monos)),
    56     REPEAT_DETERM (FIRSTGOAL (resolve_tac monos)),
    48     REPEAT_DETERM (rtac impI 1 THEN (atac 1 ORELSE tac))]
    57     REPEAT_DETERM (rtac impI 1 THEN (atac 1 ORELSE tac))]
    49 end
    58 end
    50 
    59 
    51 (* 
       
    52  proves F[f t] from F[t] where F[t] is the given theorem  
       
    53   
       
    54   - F needs to be monotone
       
    55   - f returns either SOME for a term it fires 
       
    56     and NONE elsewhere 
       
    57 *)
       
    58 fun map_thm ctxt f tac thm =
    60 fun map_thm ctxt f tac thm =
    59 let
    61 let
    60   val opt_goal_trm = map_term f (prop_of thm)
    62   val opt_goal_trm = map_term f (prop_of thm)
    61   fun prove goal = 
       
    62     Goal.prove ctxt [] [] goal (fn _ => map_thm_tac ctxt tac thm)
       
    63 in
    63 in
    64   case opt_goal_trm of
    64   case opt_goal_trm of
    65     NONE => thm
    65     NONE => thm
    66   | SOME goal => prove goal
    66   | SOME goal =>
       
    67      Goal.prove ctxt [] [] goal (fn _ => map_thm_tac ctxt tac thm) 
    67 end
    68 end
    68 
    69 
       
    70 (*
       
    71  inductive premises can be of the form
       
    72  R ... /\ P ...; split_conj picks out
       
    73  the part P ...
       
    74 *)
    69 fun transform_prem ctxt names thm =
    75 fun transform_prem ctxt names thm =
    70 let
    76 let
    71   fun split_conj names (Const ("op &", _) $ p $ q) = 
    77   fun split_conj names (Const ("op &", _) $ p $ q) = 
    72       (case head_of p of
    78       (case head_of p of
    73          Const (name, _) => if name mem names then SOME q else NONE
    79          Const (name, _) => if name mem names then SOME q else NONE
    75   | split_conj _ _ = NONE;
    81   | split_conj _ _ = NONE;
    76 in
    82 in
    77   map_thm ctxt (split_conj names) (etac conjunct2 1) thm
    83   map_thm ctxt (split_conj names) (etac conjunct2 1) thm
    78 end
    84 end
    79 
    85 
    80 fun single_case_tac ctxt pred_names pi intro  = 
    86 
       
    87 (** equivariance tactics **)
       
    88 
       
    89 fun eqvt_rel_case_tac ctxt pred_names pi intro  = 
    81 let
    90 let
    82   val thy = ProofContext.theory_of ctxt
    91   val thy = ProofContext.theory_of ctxt
    83   val cpi = Thm.cterm_of thy (mk_minus pi)
    92   val cpi = Thm.cterm_of thy (mk_minus pi)
    84   val rule = Drule.instantiate' [] [SOME cpi] @{thm permute_boolE}
    93   val rule = Drule.instantiate' [] [SOME cpi] @{thm permute_boolE}
       
    94   val permute_cancel = @{thms permute_minus_cancel(2)}
    85 in
    95 in
    86   eqvt_strict_tac ctxt [] [] THEN' 
    96   eqvt_strict_tac ctxt [] [] THEN' 
    87   SUBPROOF (fn {prems, context as ctxt, ...} =>
    97   SUBPROOF (fn {prems, context, ...} =>
    88     let
    98     let
    89       val prems' = map (transform_prem ctxt pred_names) prems
    99       val prems' = map (transform_prem ctxt pred_names) prems
    90       val side_cond_tac = EVERY' 
   100       val side_cond_tac = EVERY' 
    91         [ rtac rule, 
   101         [ rtac rule, eqvt_strict_tac context permute_cancel [],
    92           eqvt_strict_tac ctxt @{thms permute_minus_cancel(2)} [],
       
    93           resolve_tac prems' ]
   102           resolve_tac prems' ]
    94     in
   103     in
    95       HEADGOAL (rtac intro THEN_ALL_NEW (resolve_tac prems' ORELSE' side_cond_tac)) 
   104       (rtac intro THEN_ALL_NEW (resolve_tac prems' ORELSE' side_cond_tac)) 1 
    96     end) ctxt
   105     end) ctxt
    97 end
   106 end
    98 
   107 
       
   108 fun eqvt_rel_tac ctxt pred_names pi induct intros =
       
   109 let
       
   110   val cases = map (eqvt_rel_case_tac ctxt pred_names pi) intros
       
   111 in
       
   112   EVERY' (rtac induct :: cases)
       
   113 end
    99 
   114 
   100 fun prepare_pred params_no pi pred =
   115 
       
   116 (** equivariance procedure *)
       
   117 
       
   118 (* sets up goal and makes sure parameters
       
   119    are untouched PROBLEM: this violates the 
       
   120    form of eqvt lemmas *)
       
   121 fun prepare_goal params_no pi pred =
   101 let
   122 let
   102   val (c, xs) = strip_comb pred;
   123   val (c, xs) = strip_comb pred;
   103   val (xs1, xs2) = chop params_no xs
   124   val (xs1, xs2) = chop params_no xs
   104 in
   125 in
   105   HOLogic.mk_imp 
   126   HOLogic.mk_imp (pred, list_comb (c, xs1 @ map (mk_perm pi) xs2))
   106     (pred, list_comb (c, xs1 @ map (mk_perm pi) xs2))
       
   107 end
   127 end
   108 
   128 
   109 
   129 (* stores thm under name.eqvt and adds [eqvt]-attribute *)
   110 fun note_named_thm (name, thm) ctxt = 
   130 fun note_named_thm (name, thm) ctxt = 
   111 let
   131 let
   112   val thm_name = Binding.qualified_name 
   132   val thm_name = Binding.qualified_name 
   113     (Long_Name.qualify (Long_Name.base_name name) "eqvt")
   133     (Long_Name.qualify (Long_Name.base_name name) "eqvt")
   114   val attr = Attrib.internal (K eqvt_add)
   134   val attr = Attrib.internal (K eqvt_add)
   115 in
   135 in
   116   Local_Theory.note ((thm_name, [attr]), [thm]) ctxt
   136   Local_Theory.note ((thm_name, [attr]), [thm]) ctxt
   117 end
   137 end
   118 
   138 
   119 
   139 fun equivariance pred_name ctxt = 
   120 fun eqvt_rel_tac pred_name ctxt = 
       
   121 let
   140 let
   122   val thy = ProofContext.theory_of ctxt
   141   val thy = ProofContext.theory_of ctxt
   123   val ({names, ...}, {raw_induct, intrs, ...}) =
   142   val ({names, ...}, {raw_induct, intrs, ...}) =
   124     Inductive.the_inductive ctxt (Sign.intern_const thy pred_name)
   143     Inductive.the_inductive ctxt (Sign.intern_const thy pred_name)
   125   val raw_induct = atomize_induct ctxt raw_induct;
   144   val raw_induct = atomize_induct ctxt raw_induct
   126   val intros = map atomize_intr intrs;
   145   val intros = map atomize_intr intrs
   127   val params_no = length (Inductive.params_of raw_induct)
   146   val params_no = length (Inductive.params_of raw_induct)
   128   val (([raw_concl], [raw_pi]), ctxt') = 
   147   val (([raw_concl], [raw_pi]), ctxt') = ctxt 
   129     ctxt |> Variable.import_terms false [concl_of raw_induct] 
   148          |> Variable.import_terms false [concl_of raw_induct] 
   130          ||>> Variable.variant_fixes ["pi"]
   149          ||>> Variable.variant_fixes ["pi"]
   131   val pi = Free (raw_pi, @{typ perm})
   150   val pi = Free (raw_pi, @{typ perm})
   132   val preds = map (fst o HOLogic.dest_imp)
   151   val preds = map (fst o HOLogic.dest_imp)
   133     (HOLogic.dest_conj (HOLogic.dest_Trueprop raw_concl));
   152     (HOLogic.dest_conj (HOLogic.dest_Trueprop raw_concl));
   134   val goal = HOLogic.mk_Trueprop 
   153   val goal = HOLogic.mk_Trueprop 
   135     (foldr1 HOLogic.mk_conj (map (prepare_pred params_no pi) preds))
   154     (foldr1 HOLogic.mk_conj (map (prepare_goal params_no pi) preds))
   136   val thm = Goal.prove ctxt' [] [] goal (fn {context,...} => 
   155   val thms = Datatype_Aux.split_conj_thm (Goal.prove ctxt' [] [] goal 
   137     HEADGOAL (EVERY' (rtac raw_induct :: map (single_case_tac context names pi) intros)))
   156     (fn {context,...} => eqvt_rel_tac context names pi raw_induct intros 1)
   138     |> singleton (ProofContext.export ctxt' ctxt)
   157     |> singleton (ProofContext.export ctxt' ctxt))
   139   val thms = map (fn th => zero_var_indexes (th RS mp)) (Datatype_Aux.split_conj_thm thm)
   158   val thms' = map (fn th => zero_var_indexes (th RS mp)) thms
   140 in
   159 in
   141    ctxt |> fold_map note_named_thm (names ~~ thms)
   160   ctxt |> fold_map note_named_thm (names ~~ thms') |> snd  
   142         |> snd  
       
   143 end
   161 end
   144 
   162 
   145 
   163 
   146 local structure P = OuterParse and K = OuterKeyword in
   164 local structure P = OuterParse and K = OuterKeyword in
   147 
   165 
   148 val _ =
   166 val _ =
   149   OuterSyntax.local_theory "equivariance"
   167   OuterSyntax.local_theory "equivariance"
   150     "prove equivariance for inductive predicate involving nominal datatypes" K.thy_decl
   168     "prove equivariance for inductive predicate involving nominal datatypes" 
   151     (P.xname >> eqvt_rel_tac);
   169       K.thy_decl (P.xname >> equivariance);
   152 
       
   153 end;
   170 end;
   154 
   171 
   155 end (* structure *)
   172 end (* structure *)