Nominal/activities/tphols09/IDW/SB-Method.thy
changeset 415 f1be8028a4a9
equal deleted inserted replaced
414:05e5d68c9627 415:f1be8028a4a9
       
     1 theory Method
       
     2 imports Main
       
     3 begin
       
     4 
       
     5 
       
     6 
       
     7 section {* A simple method *}
       
     8 
       
     9 ML {*
       
    10 val q1_syntax = Scan.succeed
       
    11 *}
       
    12 
       
    13 ML {*
       
    14 fun q1_method ctxt = SIMPLE_METHOD'
       
    15   (Simplifier.simp_tac (Simplifier.context ctxt HOL_ss))
       
    16 *}
       
    17 
       
    18 method_setup Q1 = {* q1_syntax q1_method *} "The Q method"
       
    19 
       
    20 lemma True by Q1
       
    21 
       
    22 
       
    23 
       
    24 
       
    25 
       
    26 section {* Select a Subgoal *}
       
    27 
       
    28 ML {*
       
    29 val q2_syntax = OuterParse.nat
       
    30 *}
       
    31 
       
    32 ML {*
       
    33 fun q2_method i ctxt = SIMPLE_METHOD
       
    34   (Simplifier.simp_tac (Simplifier.context ctxt HOL_ss) i)
       
    35 *}
       
    36 
       
    37 method_setup Q2 = {* Scan.lift q2_syntax >> q2_method *}
       
    38   "The Q method"
       
    39 
       
    40 lemma "True = (True \<and> True)"
       
    41   apply (intro iffI conjI)
       
    42   apply (Q2 3)
       
    43   sorry
       
    44 
       
    45 
       
    46 
       
    47 
       
    48 
       
    49 section {* Adding Theorems *}
       
    50 
       
    51 ML {*
       
    52 val q3_syntax =
       
    53   Scan.lift (Args.add -- Args.colon) |-- Attrib.thms
       
    54 *}
       
    55 
       
    56 ML {*
       
    57 fun q3_method thms _ = SIMPLE_METHOD'
       
    58   (Tactic.resolve_tac thms)
       
    59 *}
       
    60 
       
    61 method_setup Q3 = {* q3_syntax >> q3_method *}
       
    62   "The Q method"
       
    63 
       
    64 lemma "t = t" by (Q3 add: refl)
       
    65 
       
    66 
       
    67 
       
    68 
       
    69 
       
    70 section {* Using the Context *}
       
    71 
       
    72 ML {*
       
    73 structure Q4_Rules = NamedThmsFun
       
    74 (
       
    75   val name = "q4"
       
    76   val description = "q4 rules"
       
    77 )
       
    78 *}
       
    79 setup Q4_Rules.setup
       
    80 
       
    81 ML {*
       
    82 fun q4_method ctxt = SIMPLE_METHOD'
       
    83   (Tactic.resolve_tac (Q4_Rules.get ctxt))
       
    84 *}
       
    85 
       
    86 ML {*
       
    87 Q4_Rules.add
       
    88 *}
       
    89 
       
    90 ML {*
       
    91 val q4_syntax = Method.sections [
       
    92   Args.add -- Args.colon >> K (I, Q4_Rules.add)]
       
    93 *}
       
    94 
       
    95 method_setup Q4 = {* q4_syntax >> K q4_method *}
       
    96   "The Q method"
       
    97 
       
    98 lemma "t = t" by (Q4 add: refl)
       
    99 lemma "t = t" by (Q4 add: refl sym)
       
   100 lemma "t = t" by (Q4 add: refl add: sym)
       
   101 
       
   102 lemma "t = t" using refl[q4] by Q4
       
   103 
       
   104 
       
   105 
       
   106 
       
   107 
       
   108 section {* A Method Option *}
       
   109 
       
   110 ML {*
       
   111 val q5_opts = Scan.optional
       
   112   (Args.parens (Args.$$$ "strict" >> K true))
       
   113   false
       
   114 
       
   115 val q5_syntax =
       
   116   Scan.lift q5_opts --
       
   117   (Method.sections [
       
   118     Args.add -- Args.colon >> K (I, Q4_Rules.add)] >> flat)
       
   119 *}
       
   120 
       
   121 ML {*
       
   122 fun q5_method (strict, thms) ctxt =
       
   123   let val ths = if strict then thms else Q4_Rules.get ctxt
       
   124   in SIMPLE_METHOD' (Tactic.resolve_tac ths) end
       
   125 *}
       
   126 
       
   127 method_setup Q5 = {* q5_syntax >> q5_method *}
       
   128   "The Q method"
       
   129 
       
   130 lemma "t = t" by (Q5 add: refl)
       
   131 lemma "t = t" using refl[q4] by Q5
       
   132 (*
       
   133 lemma "t = t" using refl[q4] by (Q5 (strict))
       
   134 lemma "t = t" using refl[q4] by (Q5 (strict) add: conjI)
       
   135 *)
       
   136 lemma "t = t" by (Q5 (strict) add: refl)
       
   137 
       
   138 
       
   139 end