src/HOL/Library/Code_Abstract_Nat.thy
changeset 51113 222fb6cb2c3e
child 55415 05f5fdb8d093
equal deleted inserted replaced
51112:da97167e03f7 51113:222fb6cb2c3e
       
     1 (*  Title:      HOL/Library/Code_Abstract_Nat.thy
       
     2     Author:     Stefan Berghofer, Florian Haftmann, TU Muenchen
       
     3 *)
       
     4 
       
     5 header {* Avoidance of pattern matching on natural numbers *}
       
     6 
       
     7 theory Code_Abstract_Nat
       
     8 imports Main
       
     9 begin
       
    10 
       
    11 text {*
       
    12   When natural numbers are implemented in another than the
       
    13   conventional inductive @{term "0::nat"}/@{term Suc} representation,
       
    14   it is necessary to avoid all pattern matching on natural numbers
       
    15   altogether.  This is accomplished by this theory (up to a certain
       
    16   extent).
       
    17 *}
       
    18 
       
    19 subsection {* Case analysis *}
       
    20 
       
    21 text {*
       
    22   Case analysis on natural numbers is rephrased using a conditional
       
    23   expression:
       
    24 *}
       
    25 
       
    26 lemma [code, code_unfold]:
       
    27   "nat_case = (\<lambda>f g n. if n = 0 then f else g (n - 1))"
       
    28   by (auto simp add: fun_eq_iff dest!: gr0_implies_Suc)
       
    29 
       
    30 
       
    31 subsection {* Preprocessors *}
       
    32 
       
    33 text {*
       
    34   The term @{term "Suc n"} is no longer a valid pattern.  Therefore,
       
    35   all occurrences of this term in a position where a pattern is
       
    36   expected (i.e.~on the left-hand side of a code equation) must be
       
    37   eliminated.  This can be accomplished – as far as possible – by
       
    38   applying the following transformation rule:
       
    39 *}
       
    40 
       
    41 lemma Suc_if_eq: "(\<And>n. f (Suc n) \<equiv> h n) \<Longrightarrow> f 0 \<equiv> g \<Longrightarrow>
       
    42   f n \<equiv> if n = 0 then g else h (n - 1)"
       
    43   by (rule eq_reflection) (cases n, simp_all)
       
    44 
       
    45 text {*
       
    46   The rule above is built into a preprocessor that is plugged into
       
    47   the code generator.
       
    48 *}
       
    49 
       
    50 setup {*
       
    51 let
       
    52 
       
    53 fun remove_suc thy thms =
       
    54   let
       
    55     val vname = singleton (Name.variant_list (map fst
       
    56       (fold (Term.add_var_names o Thm.full_prop_of) thms []))) "n";
       
    57     val cv = cterm_of thy (Var ((vname, 0), HOLogic.natT));
       
    58     fun lhs_of th = snd (Thm.dest_comb
       
    59       (fst (Thm.dest_comb (cprop_of th))));
       
    60     fun rhs_of th = snd (Thm.dest_comb (cprop_of th));
       
    61     fun find_vars ct = (case term_of ct of
       
    62         (Const (@{const_name Suc}, _) $ Var _) => [(cv, snd (Thm.dest_comb ct))]
       
    63       | _ $ _ =>
       
    64         let val (ct1, ct2) = Thm.dest_comb ct
       
    65         in 
       
    66           map (apfst (fn ct => Thm.apply ct ct2)) (find_vars ct1) @
       
    67           map (apfst (Thm.apply ct1)) (find_vars ct2)
       
    68         end
       
    69       | _ => []);
       
    70     val eqs = maps
       
    71       (fn th => map (pair th) (find_vars (lhs_of th))) thms;
       
    72     fun mk_thms (th, (ct, cv')) =
       
    73       let
       
    74         val th' =
       
    75           Thm.implies_elim
       
    76            (Conv.fconv_rule (Thm.beta_conversion true)
       
    77              (Drule.instantiate'
       
    78                [SOME (ctyp_of_term ct)] [SOME (Thm.lambda cv ct),
       
    79                  SOME (Thm.lambda cv' (rhs_of th)), NONE, SOME cv']
       
    80                @{thm Suc_if_eq})) (Thm.forall_intr cv' th)
       
    81       in
       
    82         case map_filter (fn th'' =>
       
    83             SOME (th'', singleton
       
    84               (Variable.trade (K (fn [th'''] => [th''' RS th']))
       
    85                 (Variable.global_thm_context th'')) th'')
       
    86           handle THM _ => NONE) thms of
       
    87             [] => NONE
       
    88           | thps =>
       
    89               let val (ths1, ths2) = split_list thps
       
    90               in SOME (subtract Thm.eq_thm (th :: ths1) thms @ ths2) end
       
    91       end
       
    92   in get_first mk_thms eqs end;
       
    93 
       
    94 fun eqn_suc_base_preproc thy thms =
       
    95   let
       
    96     val dest = fst o Logic.dest_equals o prop_of;
       
    97     val contains_suc = exists_Const (fn (c, _) => c = @{const_name Suc});
       
    98   in
       
    99     if forall (can dest) thms andalso exists (contains_suc o dest) thms
       
   100       then thms |> perhaps_loop (remove_suc thy) |> (Option.map o map) Drule.zero_var_indexes
       
   101        else NONE
       
   102   end;
       
   103 
       
   104 val eqn_suc_preproc = Code_Preproc.simple_functrans eqn_suc_base_preproc;
       
   105 
       
   106 in
       
   107 
       
   108   Code_Preproc.add_functrans ("eqn_Suc", eqn_suc_preproc)
       
   109 
       
   110 end;
       
   111 *}
       
   112 
       
   113 end