Nominal/activities/tphols09/IDW/CU-Ex1.thy
changeset 415 f1be8028a4a9
equal deleted inserted replaced
414:05e5d68c9627 415:f1be8028a4a9
       
     1 theory Ex
       
     2 imports Main
       
     3 begin
       
     4 
       
     5 text {*
       
     6   An example of an apply-proof and the
       
     7   corresponding ML-code.
       
     8 *}
       
     9 
       
    10 lemma disj_swap: 
       
    11   shows "P \<or> Q \<Longrightarrow> Q \<or> P"
       
    12 apply(erule disjE)
       
    13 apply(rule disjI2)
       
    14 apply(assumption)
       
    15 apply(rule disjI1)
       
    16 apply(assumption)
       
    17 done
       
    18 
       
    19 ML {*
       
    20 let
       
    21   val ctxt  = @{context}
       
    22   val goal  = @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"}
       
    23   val facts = []
       
    24   val schms  = ["P", "Q"]
       
    25 in
       
    26   Goal.prove ctxt schms facts goal
       
    27    (fn _ =>
       
    28       etac @{thm disjE} 1
       
    29       THEN rtac @{thm disjI2} 1
       
    30       THEN atac 1
       
    31       THEN rtac @{thm disjI1} 1
       
    32       THEN atac 1)
       
    33 end
       
    34 *}
       
    35 
       
    36 text {*  
       
    37   Tactics coded in ML can be applied / tried out in 
       
    38   apply-scripts using "tactic". "THEN" is a tactic
       
    39   combinator that just strings tactics together.
       
    40   Tactic combinators are also called tacticals.
       
    41 *}
       
    42 
       
    43 ML {*
       
    44 val foo_tac =
       
    45       (etac @{thm disjE} 1
       
    46        THEN rtac @{thm disjI2} 1
       
    47        THEN atac 1
       
    48        THEN rtac @{thm disjI1} 1
       
    49        THEN atac 1)
       
    50 *}
       
    51 
       
    52 lemma 
       
    53   shows "P \<or> Q \<Longrightarrow> Q \<or> P"
       
    54 apply(tactic {* foo_tac *})
       
    55 done
       
    56 
       
    57 text {* 
       
    58   Coding tactics with explicit subgoal addressing
       
    59   is quite brittle and limits the usage of the tactics.
       
    60   This can be avoided using the `primed' versions
       
    61   of the tactic comibinators.
       
    62 *}
       
    63 
       
    64 ML {*
       
    65 val foo_tac' =
       
    66       (etac @{thm disjE}
       
    67        THEN' rtac @{thm disjI2}
       
    68        THEN' atac
       
    69        THEN' rtac @{thm disjI1}
       
    70        THEN' atac)
       
    71 *}
       
    72 
       
    73 text {*
       
    74   Now the tactic can be used on any subgoal.
       
    75 *}
       
    76 
       
    77 lemma "P1 \<or> Q1 \<Longrightarrow> Q1 \<or> P1"
       
    78   and "P2 \<or> Q2 \<Longrightarrow> Q2 \<or> P2"
       
    79 apply(tactic {* foo_tac' 2 *})
       
    80 apply(tactic {* foo_tac' 1 *})
       
    81 done
       
    82 
       
    83 text {* 
       
    84   The type of a tactic is from a theorem
       
    85   to a lazy sequence of (successor) theorems.
       
    86 
       
    87   type tactic = thm -> thm Seq.seq 
       
    88 
       
    89   The simplest tactics are no_tac (always
       
    90   fails) and all_tac (always succeds, but 
       
    91   does not make any progress.
       
    92 *}
       
    93 
       
    94 
       
    95 ML {* fun no_tac thm = Seq.empty *}
       
    96 ML {* fun all_tac thm = Seq.single thm *}
       
    97 
       
    98 text {*
       
    99   The lazy sequence can be explored using
       
   100   "back".
       
   101 *}
       
   102 
       
   103 lemma "\<lbrakk>P \<or> Q; P \<or> Q\<rbrakk> \<Longrightarrow> Q \<or> P"
       
   104 apply(tactic {* foo_tac' 1 *})
       
   105 back
       
   106 done
       
   107 
       
   108 text {* 
       
   109   It might be surprising that a goalstate of
       
   110   an "unfinished" proof is a theorem. Below
       
   111   we show the internals.
       
   112 
       
   113   In general a goalstate is of the form
       
   114 
       
   115   A1 \<dots> An \<Longrightarrow> #C
       
   116 
       
   117   where the Ai are the open subgoals and C
       
   118   is the theorem to be proved, protected by
       
   119   the constant prop. This constant is usually
       
   120   not visible.
       
   121 *}
       
   122 
       
   123 ML {*
       
   124 fun my_print_tac ctxt thm =
       
   125 let
       
   126    val _ = tracing (Syntax.string_of_term ctxt (prop_of thm))
       
   127 in
       
   128    Seq.single thm
       
   129 end
       
   130 *}
       
   131 
       
   132 notation (output) "prop" ("#_" [1000] 1000)
       
   133 
       
   134 lemma shows "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"
       
   135 apply(tactic {* my_print_tac @{context} *})
       
   136 
       
   137 apply(rule conjI)
       
   138 apply(tactic {* my_print_tac @{context} *})
       
   139 
       
   140 apply(assumption)
       
   141 apply(tactic {* my_print_tac @{context} *})
       
   142 
       
   143 apply(assumption)
       
   144 apply(tactic {* my_print_tac @{context} *})
       
   145 done
       
   146 
       
   147 notation (output) "prop" ("_" [1000] 1000)
       
   148 
       
   149 text {* 
       
   150   There are some simple tactics corresponding
       
   151   to rule, erule and drule (from apply scripts).
       
   152   The examples should be self-explanatory.
       
   153 *}
       
   154 
       
   155 lemma shows "P \<Longrightarrow> P"
       
   156 apply(tactic {* atac 1 *})
       
   157 oops
       
   158 
       
   159 lemma shows "P \<and> Q"
       
   160 apply(tactic {* resolve_tac [@{thm conjI}] 1 *})
       
   161 oops
       
   162 
       
   163 lemma shows "P \<and> Q \<Longrightarrow> False"
       
   164 apply(tactic {* eresolve_tac [@{thm conjE}] 1 *})
       
   165 oops
       
   166 
       
   167 lemma shows "False \<and> True \<Longrightarrow> False"
       
   168 apply(tactic {* dresolve_tac [@{thm conjunct2}] 1 *})
       
   169 oops
       
   170 
       
   171 text {*
       
   172   For all sorts of operations on the goalstate there
       
   173   is a corresponding tactic. The following corresponds
       
   174   to "insert".
       
   175 *}
       
   176 
       
   177 lemma shows "True = False"
       
   178 apply(tactic {* cut_facts_tac [@{thm True_def}, @{thm False_def}] 1 *})
       
   179 oops
       
   180 
       
   181 text {* 
       
   182   Because of schematic variables, theorems need often
       
   183   to be pre-instantiated.  
       
   184 *}
       
   185 
       
   186 lemma shows "\<forall>x \<in> A. P x \<Longrightarrow> Q x"
       
   187 apply(tactic {* dresolve_tac [@{thm bspec}] 1 *})
       
   188 oops
       
   189 
       
   190 ML {* @{thm disjI1} RS @{thm conjI} *}
       
   191 
       
   192 text {*
       
   193   Tactic combinators: every tactic (THEN),
       
   194   first applicable tactic (ORELSE).
       
   195 *}
       
   196 
       
   197 ML {*
       
   198 val foo_tac' = 
       
   199   EVERY' [etac @{thm disjE}, 
       
   200           rtac @{thm disjI2},
       
   201           atac, 
       
   202           rtac @{thm disjI1}, 
       
   203           atac]
       
   204 *}
       
   205 
       
   206 
       
   207 ML {*
       
   208 val select_tac = 
       
   209   FIRST' [rtac @{thm conjI}, 
       
   210           rtac @{thm impI},
       
   211           rtac @{thm notI}, 
       
   212           rtac @{thm allI}, K all_tac]
       
   213 *}
       
   214 
       
   215 text {*
       
   216   The most convenient order to apply tactics to subgoal
       
   217   is in reverse order. 
       
   218 *}
       
   219 
       
   220 lemma 
       
   221   shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow> C" and " \<forall> x. D x" and "E \<Longrightarrow> F"
       
   222 apply(tactic {* select_tac 4 *})
       
   223 apply(tactic {* select_tac 3 *})
       
   224 apply(tactic {* select_tac 2 *})
       
   225 apply(tactic {* select_tac 1 *})
       
   226 oops
       
   227 
       
   228 text {*
       
   229   The tactic select_tac can also be written using the 
       
   230   combinator TRY.
       
   231 *}
       
   232 
       
   233 ML {*
       
   234 val select_tac' = 
       
   235   TRY o FIRST' [rtac @{thm conjI}, 
       
   236                 rtac @{thm impI},
       
   237                 rtac @{thm notI}, 
       
   238                 rtac @{thm allI}]
       
   239 *}
       
   240 
       
   241 text {*
       
   242   Repeated application of tactics. 
       
   243 *}
       
   244 
       
   245 ML {*
       
   246 val repeat_sel_tac = 
       
   247   REPEAT o CHANGED o select_tac'
       
   248 *}
       
   249 
       
   250 lemma 
       
   251   shows "A \<and> (B \<and> C)" and "A \<longrightarrow> B \<longrightarrow> C" and " \<forall> x. D x" and "E \<Longrightarrow> F"
       
   252 apply(tactic {* repeat_sel_tac 4 *})
       
   253 apply(tactic {* repeat_sel_tac 3 *})
       
   254 apply(tactic {* repeat_sel_tac 2 *})
       
   255 apply(tactic {* repeat_sel_tac 1 *})
       
   256 oops
       
   257 
       
   258 text {*  
       
   259   Application of a tactics to all new subgoals.
       
   260 *}
       
   261 
       
   262 ML {*
       
   263 val repeat_all_new_sel_tac = 
       
   264   TRY o REPEAT_ALL_NEW (CHANGED o select_tac')
       
   265 *}
       
   266 
       
   267 lemma 
       
   268   shows "A \<and> (B \<and> C)" and "A \<longrightarrow> B \<longrightarrow> C" and " \<forall> x. D x" and "E \<Longrightarrow> F"
       
   269 apply(tactic {* repeat_all_new_sel_tac 4 *})
       
   270 apply(tactic {* repeat_all_new_sel_tac 3 *})
       
   271 apply(tactic {* repeat_all_new_sel_tac 2 *})
       
   272 apply(tactic {* repeat_all_new_sel_tac 1 *})
       
   273 oops
       
   274 
       
   275 text {* 
       
   276   The tactical DEPTH_SOLVED applies tactics in depth first
       
   277   search including backtracking.
       
   278 
       
   279   This can be used to implement a decision procedure for 
       
   280   propositional intuitionistic logic inspired by work of 
       
   281   Roy Dyckhoff.
       
   282 *}
       
   283 
       
   284 
       
   285 lemma impE1:
       
   286   shows "\<lbrakk>A \<longrightarrow> B; A; \<lbrakk>A; B\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
       
   287 by iprover
       
   288 
       
   289 lemma impE2:
       
   290   shows "\<lbrakk>(C \<and> D) \<longrightarrow> B; C \<longrightarrow> (D \<longrightarrow>B) \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
       
   291 by iprover
       
   292 
       
   293 lemma impE3:
       
   294   shows "\<lbrakk>(C \<or> D) \<longrightarrow> B; \<lbrakk>C \<longrightarrow> B; D \<longrightarrow> B\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
       
   295 by iprover
       
   296 
       
   297 lemma impE4:
       
   298   shows "\<lbrakk>(C \<longrightarrow> D) \<longrightarrow> B; D \<longrightarrow> B \<Longrightarrow> C \<longrightarrow> D; B \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
       
   299 by iprover
       
   300 
       
   301 lemma impE5:
       
   302   shows "\<lbrakk>(C = D) \<longrightarrow> B; (C \<longrightarrow> D) \<longrightarrow> ((D \<longrightarrow> C) \<longrightarrow> B) \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
       
   303 by iprover
       
   304 
       
   305 ML {*
       
   306 val apply_tac =
       
   307 let
       
   308   val intros = [@{thm conjI}, @{thm disjI1}, @{thm disjI2}, @{thm impI}, @{thm iffI}]
       
   309   val elims  = [@{thm FalseE}, @{thm conjE}, @{thm disjE}, @{thm iffE},
       
   310                 @{thm impE2}, @{thm impE3}, @{thm impE4}, @{thm impE5}, @{thm impE1}]
       
   311 in
       
   312   atac 
       
   313   ORELSE' resolve_tac intros
       
   314   ORELSE' eresolve_tac elims
       
   315 end
       
   316 *}
       
   317 
       
   318 lemma
       
   319   shows "((((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P) \<longrightarrow> Q) \<longrightarrow> Q"
       
   320 apply(tactic {* (DEPTH_SOLVE o apply_tac) 1 *})
       
   321 done
       
   322 
       
   323 lemma 
       
   324   shows "((A \<longrightarrow> B) \<longrightarrow> A) \<longrightarrow> A"
       
   325 apply(tactic {* (TRY o DEPTH_SOLVE o apply_tac) 1 *})
       
   326 oops
       
   327 
       
   328 lemma
       
   329   shows "(A \<and> B \<or> ((((A \<longrightarrow> False) \<longrightarrow> False) \<longrightarrow> Q) \<or> (B \<longrightarrow> Q)) \<longrightarrow> Q) \<longrightarrow> Q"
       
   330 apply(tactic {* (TRY o DEPTH_SOLVE o apply_tac) 1 *})
       
   331 oops
       
   332 
       
   333 text {* 
       
   334   The SUBPROOF tactic gives you full access to
       
   335   all components of a goal state.
       
   336 *}
       
   337 
       
   338 ML {*
       
   339 fun subproof_tac {asms, concl, prems, params, context, schematics} =
       
   340 let
       
   341   fun out f xs = commas (map (Syntax.string_of_term context o f) xs)
       
   342   val str_of_asms   = out term_of asms
       
   343   val str_of_concl  = out term_of [concl]
       
   344   val str_of_prems  = out prop_of prems
       
   345   val str_of_params = out term_of params
       
   346   val str_of_schms  = out term_of (snd schematics)
       
   347   val s = ["params: "      ^ str_of_params,
       
   348            "schematics: "  ^ str_of_schms,
       
   349            "premises: "    ^ str_of_prems,
       
   350            "assumptions: " ^ str_of_asms,
       
   351            "conclusion: "  ^ str_of_concl]
       
   352 in
       
   353    tracing (cat_lines s); no_tac
       
   354 end
       
   355 *}
       
   356 
       
   357 lemma shows "\<And>x y. A x y \<Longrightarrow> B y x \<longrightarrow> C (?z y) x"
       
   358 apply(tactic {* SUBPROOF subproof_tac @{context} 1 *})?
       
   359 apply(rule impI)
       
   360 apply(tactic {* SUBPROOF subproof_tac @{context} 1 *})?
       
   361 oops
       
   362 
       
   363 ML {*
       
   364 val my_atac = SUBPROOF (fn {prems, ...} => resolve_tac prems 1)
       
   365 *}
       
   366 
       
   367 lemma shows "\<lbrakk>B x y; A x y; C x y \<rbrakk> \<Longrightarrow> A x y"
       
   368 apply(tactic {* my_atac @{context} 1 *})
       
   369 done
       
   370 
       
   371 
       
   372 text {*
       
   373   Setting up the goal state using the function 
       
   374   Goal.prove. It expects the goal as a term, we
       
   375   have to construct manually.
       
   376 *}
       
   377 
       
   378 ML {* 
       
   379 open HOLogic
       
   380 
       
   381 fun P n = @{term "P::nat \<Rightarrow> bool"} $ (mk_number @{typ "nat"} n) 
       
   382 
       
   383 fun rhs 1 = P 1
       
   384   | rhs n = mk_conj (P n, rhs (n - 1))
       
   385 
       
   386 fun lhs 1 n = mk_imp (mk_eq (P 1, P n), rhs n)
       
   387   | lhs m n =  mk_conj (mk_imp 
       
   388                  (mk_eq (P (m - 1), P m), rhs n), lhs (m - 1) n)
       
   389 *}
       
   390 
       
   391 ML {*
       
   392 fun de_bruijn_test ctxt n =
       
   393 let 
       
   394   val i = 2*n + 1
       
   395   val goal = mk_Trueprop (mk_imp (lhs i i, rhs i))
       
   396 in
       
   397   Syntax.string_of_term ctxt goal
       
   398   |> writeln
       
   399 end;
       
   400 
       
   401 de_bruijn_test @{context} 1   
       
   402 *}
       
   403 
       
   404 ML {*
       
   405 fun de_bruijn ctxt n =
       
   406 let 
       
   407   val i = 2*n+1
       
   408   val goal = mk_Trueprop (mk_imp (lhs i i, rhs i))
       
   409 in
       
   410   Goal.prove ctxt ["P"] [] goal
       
   411    (fn _ => (DEPTH_SOLVE o apply_tac) 1)
       
   412 end 
       
   413 *}
       
   414 
       
   415 ML{* de_bruijn @{context} 1 *}
       
   416 
       
   417 text {*
       
   418   Below is the version which constructs the terms
       
   419   using fixed variables.
       
   420 *}
       
   421 
       
   422 ML {*
       
   423 fun de_bruijn ctxt n = 
       
   424 let 
       
   425   val i = 2*n+1
       
   426   val bs = replicate (i+1) "b"
       
   427   val (nbs, ctxt') = Variable.variant_fixes bs ctxt
       
   428   val fbs = map (fn z => Free (z, @{typ "bool"})) nbs
       
   429 
       
   430   fun P n = nth fbs n 
       
   431   
       
   432   fun rhs 0 = @{term "True"}
       
   433     | rhs 1 = P 1
       
   434     | rhs n = mk_conj (P n, rhs (n - 1))
       
   435 
       
   436   fun lhs 1 n = mk_imp (mk_eq (P 1, P n), rhs n)
       
   437     | lhs m n =  mk_conj (mk_imp 
       
   438                    (mk_eq (P (m - 1), P m), rhs n), lhs (m - 1) n)
       
   439 
       
   440   val goal = mk_Trueprop (mk_imp (lhs i i, rhs i))
       
   441 in
       
   442   Goal.prove ctxt' [] [] goal
       
   443    (fn _ => (DEPTH_SOLVE o apply_tac) 1)
       
   444   (*|> singleton (ProofContext.export ctxt' ctxt)*)
       
   445 end
       
   446 *}
       
   447 
       
   448 ML {* de_bruijn @{context} 1 *}
       
   449 
       
   450 
       
   451 end
       
   452 
       
   453  
       
   454 
       
   455 
       
   456 
       
   457 
       
   458 
       
   459 
       
   460 
       
   461