src/HOL/HOL.thy
author haftmann
Thu, 26 Aug 2010 12:19:49 +0200
changeset 39006 f9837065b5e8
parent 38963 6513ea67d95d
child 39019 e46e7a9cb622
permissions -rw-r--r--
prevent line breaks after Scala symbolic operators
     1 (*  Title:      HOL/HOL.thy
     2     Author:     Tobias Nipkow, Markus Wenzel, and Larry Paulson
     3 *)
     4 
     5 header {* The basis of Higher-Order Logic *}
     6 
     7 theory HOL
     8 imports Pure "~~/src/Tools/Code_Generator"
     9 uses
    10   ("Tools/hologic.ML")
    11   "~~/src/Tools/IsaPlanner/zipper.ML"
    12   "~~/src/Tools/IsaPlanner/isand.ML"
    13   "~~/src/Tools/IsaPlanner/rw_tools.ML"
    14   "~~/src/Tools/IsaPlanner/rw_inst.ML"
    15   "~~/src/Tools/intuitionistic.ML"
    16   "~~/src/Tools/project_rule.ML"
    17   "~~/src/Tools/cong_tac.ML"
    18   "~~/src/Tools/misc_legacy.ML"
    19   "~~/src/Provers/hypsubst.ML"
    20   "~~/src/Provers/splitter.ML"
    21   "~~/src/Provers/classical.ML"
    22   "~~/src/Provers/blast.ML"
    23   "~~/src/Provers/clasimp.ML"
    24   "~~/src/Tools/coherent.ML"
    25   "~~/src/Tools/eqsubst.ML"
    26   "~~/src/Provers/quantifier1.ML"
    27   ("Tools/simpdata.ML")
    28   "~~/src/Tools/random_word.ML"
    29   "~~/src/Tools/atomize_elim.ML"
    30   "~~/src/Tools/induct.ML"
    31   ("~~/src/Tools/induct_tacs.ML")
    32   ("Tools/recfun_codegen.ML")
    33 begin
    34 
    35 setup {* Intuitionistic.method_setup @{binding iprover} *}
    36 
    37 
    38 subsection {* Primitive logic *}
    39 
    40 subsubsection {* Core syntax *}
    41 
    42 classes type
    43 default_sort type
    44 setup {* Object_Logic.add_base_sort @{sort type} *}
    45 
    46 arities
    47   "fun" :: (type, type) type
    48   itself :: (type) type
    49 
    50 typedecl bool
    51 
    52 judgment
    53   Trueprop      :: "bool => prop"                   ("(_)" 5)
    54 
    55 consts
    56   True          :: bool
    57   False         :: bool
    58   Not           :: "bool => bool"                   ("~ _" [40] 40)
    59 
    60 setup Sign.root_path
    61 
    62 consts
    63   "op &"        :: "[bool, bool] => bool"           (infixr "&" 35)
    64   "op |"        :: "[bool, bool] => bool"           (infixr "|" 30)
    65   "op -->"      :: "[bool, bool] => bool"           (infixr "-->" 25)
    66 
    67   "op ="        :: "['a, 'a] => bool"               (infixl "=" 50)
    68 
    69 setup Sign.local_path
    70 
    71 consts
    72   The           :: "('a => bool) => 'a"
    73   All           :: "('a => bool) => bool"           (binder "ALL " 10)
    74   Ex            :: "('a => bool) => bool"           (binder "EX " 10)
    75   Ex1           :: "('a => bool) => bool"           (binder "EX! " 10)
    76 
    77 
    78 subsubsection {* Additional concrete syntax *}
    79 
    80 notation (output)
    81   "op ="  (infix "=" 50)
    82 
    83 abbreviation
    84   not_equal :: "['a, 'a] => bool"  (infixl "~=" 50) where
    85   "x ~= y == ~ (x = y)"
    86 
    87 notation (output)
    88   not_equal  (infix "~=" 50)
    89 
    90 notation (xsymbols)
    91   Not  ("\<not> _" [40] 40) and
    92   "op &"  (infixr "\<and>" 35) and
    93   "op |"  (infixr "\<or>" 30) and
    94   "op -->"  (infixr "\<longrightarrow>" 25) and
    95   not_equal  (infix "\<noteq>" 50)
    96 
    97 notation (HTML output)
    98   Not  ("\<not> _" [40] 40) and
    99   "op &"  (infixr "\<and>" 35) and
   100   "op |"  (infixr "\<or>" 30) and
   101   not_equal  (infix "\<noteq>" 50)
   102 
   103 abbreviation (iff)
   104   iff :: "[bool, bool] => bool"  (infixr "<->" 25) where
   105   "A <-> B == A = B"
   106 
   107 notation (xsymbols)
   108   iff  (infixr "\<longleftrightarrow>" 25)
   109 
   110 nonterminals
   111   letbinds  letbind
   112   case_syn  cases_syn
   113 
   114 syntax
   115   "_The"        :: "[pttrn, bool] => 'a"                 ("(3THE _./ _)" [0, 10] 10)
   116 
   117   "_bind"       :: "[pttrn, 'a] => letbind"              ("(2_ =/ _)" 10)
   118   ""            :: "letbind => letbinds"                 ("_")
   119   "_binds"      :: "[letbind, letbinds] => letbinds"     ("_;/ _")
   120   "_Let"        :: "[letbinds, 'a] => 'a"                ("(let (_)/ in (_))" [0, 10] 10)
   121 
   122   "_case_syntax":: "['a, cases_syn] => 'b"               ("(case _ of/ _)" 10)
   123   "_case1"      :: "['a, 'b] => case_syn"                ("(2_ =>/ _)" 10)
   124   ""            :: "case_syn => cases_syn"               ("_")
   125   "_case2"      :: "[case_syn, cases_syn] => cases_syn"  ("_/ | _")
   126 
   127 translations
   128   "THE x. P"              == "CONST The (%x. P)"
   129 
   130 print_translation {*
   131   [(@{const_syntax The}, fn [Abs abs] =>
   132       let val (x, t) = atomic_abs_tr' abs
   133       in Syntax.const @{syntax_const "_The"} $ x $ t end)]
   134 *}  -- {* To avoid eta-contraction of body *}
   135 
   136 syntax (xsymbols)
   137   "_case1"      :: "['a, 'b] => case_syn"                ("(2_ \<Rightarrow>/ _)" 10)
   138 
   139 notation (xsymbols)
   140   All  (binder "\<forall>" 10) and
   141   Ex  (binder "\<exists>" 10) and
   142   Ex1  (binder "\<exists>!" 10)
   143 
   144 notation (HTML output)
   145   All  (binder "\<forall>" 10) and
   146   Ex  (binder "\<exists>" 10) and
   147   Ex1  (binder "\<exists>!" 10)
   148 
   149 notation (HOL)
   150   All  (binder "! " 10) and
   151   Ex  (binder "? " 10) and
   152   Ex1  (binder "?! " 10)
   153 
   154 
   155 subsubsection {* Axioms and basic definitions *}
   156 
   157 axioms
   158   refl:           "t = (t::'a)"
   159   subst:          "s = t \<Longrightarrow> P s \<Longrightarrow> P t"
   160   ext:            "(!!x::'a. (f x ::'b) = g x) ==> (%x. f x) = (%x. g x)"
   161     -- {*Extensionality is built into the meta-logic, and this rule expresses
   162          a related property.  It is an eta-expanded version of the traditional
   163          rule, and similar to the ABS rule of HOL*}
   164 
   165   the_eq_trivial: "(THE x. x = a) = (a::'a)"
   166 
   167   impI:           "(P ==> Q) ==> P-->Q"
   168   mp:             "[| P-->Q;  P |] ==> Q"
   169 
   170 
   171 defs
   172   True_def:     "True      == ((%x::bool. x) = (%x. x))"
   173   All_def:      "All(P)    == (P = (%x. True))"
   174   Ex_def:       "Ex(P)     == !Q. (!x. P x --> Q) --> Q"
   175   False_def:    "False     == (!P. P)"
   176   not_def:      "~ P       == P-->False"
   177   and_def:      "P & Q     == !R. (P-->Q-->R) --> R"
   178   or_def:       "P | Q     == !R. (P-->R) --> (Q-->R) --> R"
   179   Ex1_def:      "Ex1(P)    == ? x. P(x) & (! y. P(y) --> y=x)"
   180 
   181 axioms
   182   iff:          "(P-->Q) --> (Q-->P) --> (P=Q)"
   183   True_or_False:  "(P=True) | (P=False)"
   184 
   185 finalconsts
   186   "op ="
   187   "op -->"
   188   The
   189 
   190 definition If :: "bool \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("(if (_)/ then (_)/ else (_))" [0, 0, 10] 10) where
   191   "If P x y \<equiv> (THE z::'a. (P=True --> z=x) & (P=False --> z=y))"
   192 
   193 definition Let :: "'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b" where
   194   "Let s f \<equiv> f s"
   195 
   196 translations
   197   "_Let (_binds b bs) e"  == "_Let b (_Let bs e)"
   198   "let x = a in e"        == "CONST Let a (%x. e)"
   199 
   200 axiomatization
   201   undefined :: 'a
   202 
   203 class default =
   204   fixes default :: 'a
   205 
   206 
   207 subsection {* Fundamental rules *}
   208 
   209 subsubsection {* Equality *}
   210 
   211 lemma sym: "s = t ==> t = s"
   212   by (erule subst) (rule refl)
   213 
   214 lemma ssubst: "t = s ==> P s ==> P t"
   215   by (drule sym) (erule subst)
   216 
   217 lemma trans: "[| r=s; s=t |] ==> r=t"
   218   by (erule subst)
   219 
   220 lemma meta_eq_to_obj_eq: 
   221   assumes meq: "A == B"
   222   shows "A = B"
   223   by (unfold meq) (rule refl)
   224 
   225 text {* Useful with @{text erule} for proving equalities from known equalities. *}
   226      (* a = b
   227         |   |
   228         c = d   *)
   229 lemma box_equals: "[| a=b;  a=c;  b=d |] ==> c=d"
   230 apply (rule trans)
   231 apply (rule trans)
   232 apply (rule sym)
   233 apply assumption+
   234 done
   235 
   236 text {* For calculational reasoning: *}
   237 
   238 lemma forw_subst: "a = b ==> P b ==> P a"
   239   by (rule ssubst)
   240 
   241 lemma back_subst: "P a ==> a = b ==> P b"
   242   by (rule subst)
   243 
   244 
   245 subsubsection {* Congruence rules for application *}
   246 
   247 text {* Similar to @{text AP_THM} in Gordon's HOL. *}
   248 lemma fun_cong: "(f::'a=>'b) = g ==> f(x)=g(x)"
   249 apply (erule subst)
   250 apply (rule refl)
   251 done
   252 
   253 text {* Similar to @{text AP_TERM} in Gordon's HOL and FOL's @{text subst_context}. *}
   254 lemma arg_cong: "x=y ==> f(x)=f(y)"
   255 apply (erule subst)
   256 apply (rule refl)
   257 done
   258 
   259 lemma arg_cong2: "\<lbrakk> a = b; c = d \<rbrakk> \<Longrightarrow> f a c = f b d"
   260 apply (erule ssubst)+
   261 apply (rule refl)
   262 done
   263 
   264 lemma cong: "[| f = g; (x::'a) = y |] ==> f x = g y"
   265 apply (erule subst)+
   266 apply (rule refl)
   267 done
   268 
   269 ML {* val cong_tac = Cong_Tac.cong_tac @{thm cong} *}
   270 
   271 
   272 subsubsection {* Equality of booleans -- iff *}
   273 
   274 lemma iffI: assumes "P ==> Q" and "Q ==> P" shows "P=Q"
   275   by (iprover intro: iff [THEN mp, THEN mp] impI assms)
   276 
   277 lemma iffD2: "[| P=Q; Q |] ==> P"
   278   by (erule ssubst)
   279 
   280 lemma rev_iffD2: "[| Q; P=Q |] ==> P"
   281   by (erule iffD2)
   282 
   283 lemma iffD1: "Q = P \<Longrightarrow> Q \<Longrightarrow> P"
   284   by (drule sym) (rule iffD2)
   285 
   286 lemma rev_iffD1: "Q \<Longrightarrow> Q = P \<Longrightarrow> P"
   287   by (drule sym) (rule rev_iffD2)
   288 
   289 lemma iffE:
   290   assumes major: "P=Q"
   291     and minor: "[| P --> Q; Q --> P |] ==> R"
   292   shows R
   293   by (iprover intro: minor impI major [THEN iffD2] major [THEN iffD1])
   294 
   295 
   296 subsubsection {*True*}
   297 
   298 lemma TrueI: "True"
   299   unfolding True_def by (rule refl)
   300 
   301 lemma eqTrueI: "P ==> P = True"
   302   by (iprover intro: iffI TrueI)
   303 
   304 lemma eqTrueE: "P = True ==> P"
   305   by (erule iffD2) (rule TrueI)
   306 
   307 
   308 subsubsection {*Universal quantifier*}
   309 
   310 lemma allI: assumes "!!x::'a. P(x)" shows "ALL x. P(x)"
   311   unfolding All_def by (iprover intro: ext eqTrueI assms)
   312 
   313 lemma spec: "ALL x::'a. P(x) ==> P(x)"
   314 apply (unfold All_def)
   315 apply (rule eqTrueE)
   316 apply (erule fun_cong)
   317 done
   318 
   319 lemma allE:
   320   assumes major: "ALL x. P(x)"
   321     and minor: "P(x) ==> R"
   322   shows R
   323   by (iprover intro: minor major [THEN spec])
   324 
   325 lemma all_dupE:
   326   assumes major: "ALL x. P(x)"
   327     and minor: "[| P(x); ALL x. P(x) |] ==> R"
   328   shows R
   329   by (iprover intro: minor major major [THEN spec])
   330 
   331 
   332 subsubsection {* False *}
   333 
   334 text {*
   335   Depends upon @{text spec}; it is impossible to do propositional
   336   logic before quantifiers!
   337 *}
   338 
   339 lemma FalseE: "False ==> P"
   340   apply (unfold False_def)
   341   apply (erule spec)
   342   done
   343 
   344 lemma False_neq_True: "False = True ==> P"
   345   by (erule eqTrueE [THEN FalseE])
   346 
   347 
   348 subsubsection {* Negation *}
   349 
   350 lemma notI:
   351   assumes "P ==> False"
   352   shows "~P"
   353   apply (unfold not_def)
   354   apply (iprover intro: impI assms)
   355   done
   356 
   357 lemma False_not_True: "False ~= True"
   358   apply (rule notI)
   359   apply (erule False_neq_True)
   360   done
   361 
   362 lemma True_not_False: "True ~= False"
   363   apply (rule notI)
   364   apply (drule sym)
   365   apply (erule False_neq_True)
   366   done
   367 
   368 lemma notE: "[| ~P;  P |] ==> R"
   369   apply (unfold not_def)
   370   apply (erule mp [THEN FalseE])
   371   apply assumption
   372   done
   373 
   374 lemma notI2: "(P \<Longrightarrow> \<not> Pa) \<Longrightarrow> (P \<Longrightarrow> Pa) \<Longrightarrow> \<not> P"
   375   by (erule notE [THEN notI]) (erule meta_mp)
   376 
   377 
   378 subsubsection {*Implication*}
   379 
   380 lemma impE:
   381   assumes "P-->Q" "P" "Q ==> R"
   382   shows "R"
   383 by (iprover intro: assms mp)
   384 
   385 (* Reduces Q to P-->Q, allowing substitution in P. *)
   386 lemma rev_mp: "[| P;  P --> Q |] ==> Q"
   387 by (iprover intro: mp)
   388 
   389 lemma contrapos_nn:
   390   assumes major: "~Q"
   391       and minor: "P==>Q"
   392   shows "~P"
   393 by (iprover intro: notI minor major [THEN notE])
   394 
   395 (*not used at all, but we already have the other 3 combinations *)
   396 lemma contrapos_pn:
   397   assumes major: "Q"
   398       and minor: "P ==> ~Q"
   399   shows "~P"
   400 by (iprover intro: notI minor major notE)
   401 
   402 lemma not_sym: "t ~= s ==> s ~= t"
   403   by (erule contrapos_nn) (erule sym)
   404 
   405 lemma eq_neq_eq_imp_neq: "[| x = a ; a ~= b; b = y |] ==> x ~= y"
   406   by (erule subst, erule ssubst, assumption)
   407 
   408 (*still used in HOLCF*)
   409 lemma rev_contrapos:
   410   assumes pq: "P ==> Q"
   411       and nq: "~Q"
   412   shows "~P"
   413 apply (rule nq [THEN contrapos_nn])
   414 apply (erule pq)
   415 done
   416 
   417 subsubsection {*Existential quantifier*}
   418 
   419 lemma exI: "P x ==> EX x::'a. P x"
   420 apply (unfold Ex_def)
   421 apply (iprover intro: allI allE impI mp)
   422 done
   423 
   424 lemma exE:
   425   assumes major: "EX x::'a. P(x)"
   426       and minor: "!!x. P(x) ==> Q"
   427   shows "Q"
   428 apply (rule major [unfolded Ex_def, THEN spec, THEN mp])
   429 apply (iprover intro: impI [THEN allI] minor)
   430 done
   431 
   432 
   433 subsubsection {*Conjunction*}
   434 
   435 lemma conjI: "[| P; Q |] ==> P&Q"
   436 apply (unfold and_def)
   437 apply (iprover intro: impI [THEN allI] mp)
   438 done
   439 
   440 lemma conjunct1: "[| P & Q |] ==> P"
   441 apply (unfold and_def)
   442 apply (iprover intro: impI dest: spec mp)
   443 done
   444 
   445 lemma conjunct2: "[| P & Q |] ==> Q"
   446 apply (unfold and_def)
   447 apply (iprover intro: impI dest: spec mp)
   448 done
   449 
   450 lemma conjE:
   451   assumes major: "P&Q"
   452       and minor: "[| P; Q |] ==> R"
   453   shows "R"
   454 apply (rule minor)
   455 apply (rule major [THEN conjunct1])
   456 apply (rule major [THEN conjunct2])
   457 done
   458 
   459 lemma context_conjI:
   460   assumes "P" "P ==> Q" shows "P & Q"
   461 by (iprover intro: conjI assms)
   462 
   463 
   464 subsubsection {*Disjunction*}
   465 
   466 lemma disjI1: "P ==> P|Q"
   467 apply (unfold or_def)
   468 apply (iprover intro: allI impI mp)
   469 done
   470 
   471 lemma disjI2: "Q ==> P|Q"
   472 apply (unfold or_def)
   473 apply (iprover intro: allI impI mp)
   474 done
   475 
   476 lemma disjE:
   477   assumes major: "P|Q"
   478       and minorP: "P ==> R"
   479       and minorQ: "Q ==> R"
   480   shows "R"
   481 by (iprover intro: minorP minorQ impI
   482                  major [unfolded or_def, THEN spec, THEN mp, THEN mp])
   483 
   484 
   485 subsubsection {*Classical logic*}
   486 
   487 lemma classical:
   488   assumes prem: "~P ==> P"
   489   shows "P"
   490 apply (rule True_or_False [THEN disjE, THEN eqTrueE])
   491 apply assumption
   492 apply (rule notI [THEN prem, THEN eqTrueI])
   493 apply (erule subst)
   494 apply assumption
   495 done
   496 
   497 lemmas ccontr = FalseE [THEN classical, standard]
   498 
   499 (*notE with premises exchanged; it discharges ~R so that it can be used to
   500   make elimination rules*)
   501 lemma rev_notE:
   502   assumes premp: "P"
   503       and premnot: "~R ==> ~P"
   504   shows "R"
   505 apply (rule ccontr)
   506 apply (erule notE [OF premnot premp])
   507 done
   508 
   509 (*Double negation law*)
   510 lemma notnotD: "~~P ==> P"
   511 apply (rule classical)
   512 apply (erule notE)
   513 apply assumption
   514 done
   515 
   516 lemma contrapos_pp:
   517   assumes p1: "Q"
   518       and p2: "~P ==> ~Q"
   519   shows "P"
   520 by (iprover intro: classical p1 p2 notE)
   521 
   522 
   523 subsubsection {*Unique existence*}
   524 
   525 lemma ex1I:
   526   assumes "P a" "!!x. P(x) ==> x=a"
   527   shows "EX! x. P(x)"
   528 by (unfold Ex1_def, iprover intro: assms exI conjI allI impI)
   529 
   530 text{*Sometimes easier to use: the premises have no shared variables.  Safe!*}
   531 lemma ex_ex1I:
   532   assumes ex_prem: "EX x. P(x)"
   533       and eq: "!!x y. [| P(x); P(y) |] ==> x=y"
   534   shows "EX! x. P(x)"
   535 by (iprover intro: ex_prem [THEN exE] ex1I eq)
   536 
   537 lemma ex1E:
   538   assumes major: "EX! x. P(x)"
   539       and minor: "!!x. [| P(x);  ALL y. P(y) --> y=x |] ==> R"
   540   shows "R"
   541 apply (rule major [unfolded Ex1_def, THEN exE])
   542 apply (erule conjE)
   543 apply (iprover intro: minor)
   544 done
   545 
   546 lemma ex1_implies_ex: "EX! x. P x ==> EX x. P x"
   547 apply (erule ex1E)
   548 apply (rule exI)
   549 apply assumption
   550 done
   551 
   552 
   553 subsubsection {*THE: definite description operator*}
   554 
   555 lemma the_equality:
   556   assumes prema: "P a"
   557       and premx: "!!x. P x ==> x=a"
   558   shows "(THE x. P x) = a"
   559 apply (rule trans [OF _ the_eq_trivial])
   560 apply (rule_tac f = "The" in arg_cong)
   561 apply (rule ext)
   562 apply (rule iffI)
   563  apply (erule premx)
   564 apply (erule ssubst, rule prema)
   565 done
   566 
   567 lemma theI:
   568   assumes "P a" and "!!x. P x ==> x=a"
   569   shows "P (THE x. P x)"
   570 by (iprover intro: assms the_equality [THEN ssubst])
   571 
   572 lemma theI': "EX! x. P x ==> P (THE x. P x)"
   573 apply (erule ex1E)
   574 apply (erule theI)
   575 apply (erule allE)
   576 apply (erule mp)
   577 apply assumption
   578 done
   579 
   580 (*Easier to apply than theI: only one occurrence of P*)
   581 lemma theI2:
   582   assumes "P a" "!!x. P x ==> x=a" "!!x. P x ==> Q x"
   583   shows "Q (THE x. P x)"
   584 by (iprover intro: assms theI)
   585 
   586 lemma the1I2: assumes "EX! x. P x" "\<And>x. P x \<Longrightarrow> Q x" shows "Q (THE x. P x)"
   587 by(iprover intro:assms(2) theI2[where P=P and Q=Q] ex1E[OF assms(1)]
   588            elim:allE impE)
   589 
   590 lemma the1_equality [elim?]: "[| EX!x. P x; P a |] ==> (THE x. P x) = a"
   591 apply (rule the_equality)
   592 apply  assumption
   593 apply (erule ex1E)
   594 apply (erule all_dupE)
   595 apply (drule mp)
   596 apply  assumption
   597 apply (erule ssubst)
   598 apply (erule allE)
   599 apply (erule mp)
   600 apply assumption
   601 done
   602 
   603 lemma the_sym_eq_trivial: "(THE y. x=y) = x"
   604 apply (rule the_equality)
   605 apply (rule refl)
   606 apply (erule sym)
   607 done
   608 
   609 
   610 subsubsection {*Classical intro rules for disjunction and existential quantifiers*}
   611 
   612 lemma disjCI:
   613   assumes "~Q ==> P" shows "P|Q"
   614 apply (rule classical)
   615 apply (iprover intro: assms disjI1 disjI2 notI elim: notE)
   616 done
   617 
   618 lemma excluded_middle: "~P | P"
   619 by (iprover intro: disjCI)
   620 
   621 text {*
   622   case distinction as a natural deduction rule.
   623   Note that @{term "~P"} is the second case, not the first
   624 *}
   625 lemma case_split [case_names True False]:
   626   assumes prem1: "P ==> Q"
   627       and prem2: "~P ==> Q"
   628   shows "Q"
   629 apply (rule excluded_middle [THEN disjE])
   630 apply (erule prem2)
   631 apply (erule prem1)
   632 done
   633 
   634 (*Classical implies (-->) elimination. *)
   635 lemma impCE:
   636   assumes major: "P-->Q"
   637       and minor: "~P ==> R" "Q ==> R"
   638   shows "R"
   639 apply (rule excluded_middle [of P, THEN disjE])
   640 apply (iprover intro: minor major [THEN mp])+
   641 done
   642 
   643 (*This version of --> elimination works on Q before P.  It works best for
   644   those cases in which P holds "almost everywhere".  Can't install as
   645   default: would break old proofs.*)
   646 lemma impCE':
   647   assumes major: "P-->Q"
   648       and minor: "Q ==> R" "~P ==> R"
   649   shows "R"
   650 apply (rule excluded_middle [of P, THEN disjE])
   651 apply (iprover intro: minor major [THEN mp])+
   652 done
   653 
   654 (*Classical <-> elimination. *)
   655 lemma iffCE:
   656   assumes major: "P=Q"
   657       and minor: "[| P; Q |] ==> R"  "[| ~P; ~Q |] ==> R"
   658   shows "R"
   659 apply (rule major [THEN iffE])
   660 apply (iprover intro: minor elim: impCE notE)
   661 done
   662 
   663 lemma exCI:
   664   assumes "ALL x. ~P(x) ==> P(a)"
   665   shows "EX x. P(x)"
   666 apply (rule ccontr)
   667 apply (iprover intro: assms exI allI notI notE [of "\<exists>x. P x"])
   668 done
   669 
   670 
   671 subsubsection {* Intuitionistic Reasoning *}
   672 
   673 lemma impE':
   674   assumes 1: "P --> Q"
   675     and 2: "Q ==> R"
   676     and 3: "P --> Q ==> P"
   677   shows R
   678 proof -
   679   from 3 and 1 have P .
   680   with 1 have Q by (rule impE)
   681   with 2 show R .
   682 qed
   683 
   684 lemma allE':
   685   assumes 1: "ALL x. P x"
   686     and 2: "P x ==> ALL x. P x ==> Q"
   687   shows Q
   688 proof -
   689   from 1 have "P x" by (rule spec)
   690   from this and 1 show Q by (rule 2)
   691 qed
   692 
   693 lemma notE':
   694   assumes 1: "~ P"
   695     and 2: "~ P ==> P"
   696   shows R
   697 proof -
   698   from 2 and 1 have P .
   699   with 1 show R by (rule notE)
   700 qed
   701 
   702 lemma TrueE: "True ==> P ==> P" .
   703 lemma notFalseE: "~ False ==> P ==> P" .
   704 
   705 lemmas [Pure.elim!] = disjE iffE FalseE conjE exE TrueE notFalseE
   706   and [Pure.intro!] = iffI conjI impI TrueI notI allI refl
   707   and [Pure.elim 2] = allE notE' impE'
   708   and [Pure.intro] = exI disjI2 disjI1
   709 
   710 lemmas [trans] = trans
   711   and [sym] = sym not_sym
   712   and [Pure.elim?] = iffD1 iffD2 impE
   713 
   714 use "Tools/hologic.ML"
   715 
   716 
   717 subsubsection {* Atomizing meta-level connectives *}
   718 
   719 axiomatization where
   720   eq_reflection: "x = y \<Longrightarrow> x \<equiv> y" (*admissible axiom*)
   721 
   722 lemma atomize_all [atomize]: "(!!x. P x) == Trueprop (ALL x. P x)"
   723 proof
   724   assume "!!x. P x"
   725   then show "ALL x. P x" ..
   726 next
   727   assume "ALL x. P x"
   728   then show "!!x. P x" by (rule allE)
   729 qed
   730 
   731 lemma atomize_imp [atomize]: "(A ==> B) == Trueprop (A --> B)"
   732 proof
   733   assume r: "A ==> B"
   734   show "A --> B" by (rule impI) (rule r)
   735 next
   736   assume "A --> B" and A
   737   then show B by (rule mp)
   738 qed
   739 
   740 lemma atomize_not: "(A ==> False) == Trueprop (~A)"
   741 proof
   742   assume r: "A ==> False"
   743   show "~A" by (rule notI) (rule r)
   744 next
   745   assume "~A" and A
   746   then show False by (rule notE)
   747 qed
   748 
   749 lemma atomize_eq [atomize]: "(x == y) == Trueprop (x = y)"
   750 proof
   751   assume "x == y"
   752   show "x = y" by (unfold `x == y`) (rule refl)
   753 next
   754   assume "x = y"
   755   then show "x == y" by (rule eq_reflection)
   756 qed
   757 
   758 lemma atomize_conj [atomize]: "(A &&& B) == Trueprop (A & B)"
   759 proof
   760   assume conj: "A &&& B"
   761   show "A & B"
   762   proof (rule conjI)
   763     from conj show A by (rule conjunctionD1)
   764     from conj show B by (rule conjunctionD2)
   765   qed
   766 next
   767   assume conj: "A & B"
   768   show "A &&& B"
   769   proof -
   770     from conj show A ..
   771     from conj show B ..
   772   qed
   773 qed
   774 
   775 lemmas [symmetric, rulify] = atomize_all atomize_imp
   776   and [symmetric, defn] = atomize_all atomize_imp atomize_eq
   777 
   778 
   779 subsubsection {* Atomizing elimination rules *}
   780 
   781 setup AtomizeElim.setup
   782 
   783 lemma atomize_exL[atomize_elim]: "(!!x. P x ==> Q) == ((EX x. P x) ==> Q)"
   784   by rule iprover+
   785 
   786 lemma atomize_conjL[atomize_elim]: "(A ==> B ==> C) == (A & B ==> C)"
   787   by rule iprover+
   788 
   789 lemma atomize_disjL[atomize_elim]: "((A ==> C) ==> (B ==> C) ==> C) == ((A | B ==> C) ==> C)"
   790   by rule iprover+
   791 
   792 lemma atomize_elimL[atomize_elim]: "(!!B. (A ==> B) ==> B) == Trueprop A" ..
   793 
   794 
   795 subsection {* Package setup *}
   796 
   797 subsubsection {* Sledgehammer setup *}
   798 
   799 text {*
   800 Theorems blacklisted to Sledgehammer. These theorems typically produce clauses
   801 that are prolific (match too many equality or membership literals) and relate to
   802 seldom-used facts. Some duplicate other rules.
   803 *}
   804 
   805 ML {*
   806 structure No_ATPs = Named_Thms
   807 (
   808   val name = "no_atp"
   809   val description = "theorems that should be filtered out by Sledgehammer"
   810 )
   811 *}
   812 
   813 setup {* No_ATPs.setup *}
   814 
   815 
   816 subsubsection {* Classical Reasoner setup *}
   817 
   818 lemma imp_elim: "P --> Q ==> (~ R ==> P) ==> (Q ==> R) ==> R"
   819   by (rule classical) iprover
   820 
   821 lemma swap: "~ P ==> (~ R ==> P) ==> R"
   822   by (rule classical) iprover
   823 
   824 lemma thin_refl:
   825   "\<And>X. \<lbrakk> x=x; PROP W \<rbrakk> \<Longrightarrow> PROP W" .
   826 
   827 ML {*
   828 structure Hypsubst = HypsubstFun(
   829 struct
   830   structure Simplifier = Simplifier
   831   val dest_eq = HOLogic.dest_eq
   832   val dest_Trueprop = HOLogic.dest_Trueprop
   833   val dest_imp = HOLogic.dest_imp
   834   val eq_reflection = @{thm eq_reflection}
   835   val rev_eq_reflection = @{thm meta_eq_to_obj_eq}
   836   val imp_intr = @{thm impI}
   837   val rev_mp = @{thm rev_mp}
   838   val subst = @{thm subst}
   839   val sym = @{thm sym}
   840   val thin_refl = @{thm thin_refl};
   841   val prop_subst = @{lemma "PROP P t ==> PROP prop (x = t ==> PROP P x)"
   842                      by (unfold prop_def) (drule eq_reflection, unfold)}
   843 end);
   844 open Hypsubst;
   845 
   846 structure Classical = ClassicalFun(
   847 struct
   848   val imp_elim = @{thm imp_elim}
   849   val not_elim = @{thm notE}
   850   val swap = @{thm swap}
   851   val classical = @{thm classical}
   852   val sizef = Drule.size_of_thm
   853   val hyp_subst_tacs = [Hypsubst.hyp_subst_tac]
   854 end);
   855 
   856 structure Basic_Classical: BASIC_CLASSICAL = Classical; 
   857 open Basic_Classical;
   858 
   859 ML_Antiquote.value "claset"
   860   (Scan.succeed "Classical.claset_of (ML_Context.the_local_context ())");
   861 *}
   862 
   863 setup Classical.setup
   864 
   865 setup {*
   866 let
   867   fun non_bool_eq (@{const_name "op ="}, Type (_, [T, _])) = T <> @{typ bool}
   868     | non_bool_eq _ = false;
   869   val hyp_subst_tac' =
   870     SUBGOAL (fn (goal, i) =>
   871       if Term.exists_Const non_bool_eq goal
   872       then Hypsubst.hyp_subst_tac i
   873       else no_tac);
   874 in
   875   Hypsubst.hypsubst_setup
   876   (*prevent substitution on bool*)
   877   #> Context_Rules.addSWrapper (fn tac => hyp_subst_tac' ORELSE' tac)
   878 end
   879 *}
   880 
   881 declare iffI [intro!]
   882   and notI [intro!]
   883   and impI [intro!]
   884   and disjCI [intro!]
   885   and conjI [intro!]
   886   and TrueI [intro!]
   887   and refl [intro!]
   888 
   889 declare iffCE [elim!]
   890   and FalseE [elim!]
   891   and impCE [elim!]
   892   and disjE [elim!]
   893   and conjE [elim!]
   894 
   895 declare ex_ex1I [intro!]
   896   and allI [intro!]
   897   and the_equality [intro]
   898   and exI [intro]
   899 
   900 declare exE [elim!]
   901   allE [elim]
   902 
   903 ML {* val HOL_cs = @{claset} *}
   904 
   905 lemma contrapos_np: "~ Q ==> (~ P ==> Q) ==> P"
   906   apply (erule swap)
   907   apply (erule (1) meta_mp)
   908   done
   909 
   910 declare ex_ex1I [rule del, intro! 2]
   911   and ex1I [intro]
   912 
   913 lemmas [intro?] = ext
   914   and [elim?] = ex1_implies_ex
   915 
   916 (*Better then ex1E for classical reasoner: needs no quantifier duplication!*)
   917 lemma alt_ex1E [elim!]:
   918   assumes major: "\<exists>!x. P x"
   919       and prem: "\<And>x. \<lbrakk> P x; \<forall>y y'. P y \<and> P y' \<longrightarrow> y = y' \<rbrakk> \<Longrightarrow> R"
   920   shows R
   921 apply (rule ex1E [OF major])
   922 apply (rule prem)
   923 apply (tactic {* ares_tac @{thms allI} 1 *})+
   924 apply (tactic {* etac (Classical.dup_elim @{thm allE}) 1 *})
   925 apply iprover
   926 done
   927 
   928 ML {*
   929 structure Blast = Blast
   930 (
   931   val thy = @{theory}
   932   type claset = Classical.claset
   933   val equality_name = @{const_name "op ="}
   934   val not_name = @{const_name Not}
   935   val notE = @{thm notE}
   936   val ccontr = @{thm ccontr}
   937   val contr_tac = Classical.contr_tac
   938   val dup_intr = Classical.dup_intr
   939   val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac
   940   val rep_cs = Classical.rep_cs
   941   val cla_modifiers = Classical.cla_modifiers
   942   val cla_meth' = Classical.cla_meth'
   943 );
   944 val blast_tac = Blast.blast_tac;
   945 *}
   946 
   947 setup Blast.setup
   948 
   949 
   950 subsubsection {* Simplifier *}
   951 
   952 lemma eta_contract_eq: "(%s. f s) = f" ..
   953 
   954 lemma simp_thms:
   955   shows not_not: "(~ ~ P) = P"
   956   and Not_eq_iff: "((~P) = (~Q)) = (P = Q)"
   957   and
   958     "(P ~= Q) = (P = (~Q))"
   959     "(P | ~P) = True"    "(~P | P) = True"
   960     "(x = x) = True"
   961   and not_True_eq_False [code]: "(\<not> True) = False"
   962   and not_False_eq_True [code]: "(\<not> False) = True"
   963   and
   964     "(~P) ~= P"  "P ~= (~P)"
   965     "(True=P) = P"
   966   and eq_True: "(P = True) = P"
   967   and "(False=P) = (~P)"
   968   and eq_False: "(P = False) = (\<not> P)"
   969   and
   970     "(True --> P) = P"  "(False --> P) = True"
   971     "(P --> True) = True"  "(P --> P) = True"
   972     "(P --> False) = (~P)"  "(P --> ~P) = (~P)"
   973     "(P & True) = P"  "(True & P) = P"
   974     "(P & False) = False"  "(False & P) = False"
   975     "(P & P) = P"  "(P & (P & Q)) = (P & Q)"
   976     "(P & ~P) = False"    "(~P & P) = False"
   977     "(P | True) = True"  "(True | P) = True"
   978     "(P | False) = P"  "(False | P) = P"
   979     "(P | P) = P"  "(P | (P | Q)) = (P | Q)" and
   980     "(ALL x. P) = P"  "(EX x. P) = P"  "EX x. x=t"  "EX x. t=x"
   981   and
   982     "!!P. (EX x. x=t & P(x)) = P(t)"
   983     "!!P. (EX x. t=x & P(x)) = P(t)"
   984     "!!P. (ALL x. x=t --> P(x)) = P(t)"
   985     "!!P. (ALL x. t=x --> P(x)) = P(t)"
   986   by (blast, blast, blast, blast, blast, iprover+)
   987 
   988 lemma disj_absorb: "(A | A) = A"
   989   by blast
   990 
   991 lemma disj_left_absorb: "(A | (A | B)) = (A | B)"
   992   by blast
   993 
   994 lemma conj_absorb: "(A & A) = A"
   995   by blast
   996 
   997 lemma conj_left_absorb: "(A & (A & B)) = (A & B)"
   998   by blast
   999 
  1000 lemma eq_ac:
  1001   shows eq_commute: "(a=b) = (b=a)"
  1002     and eq_left_commute: "(P=(Q=R)) = (Q=(P=R))"
  1003     and eq_assoc: "((P=Q)=R) = (P=(Q=R))" by (iprover, blast+)
  1004 lemma neq_commute: "(a~=b) = (b~=a)" by iprover
  1005 
  1006 lemma conj_comms:
  1007   shows conj_commute: "(P&Q) = (Q&P)"
  1008     and conj_left_commute: "(P&(Q&R)) = (Q&(P&R))" by iprover+
  1009 lemma conj_assoc: "((P&Q)&R) = (P&(Q&R))" by iprover
  1010 
  1011 lemmas conj_ac = conj_commute conj_left_commute conj_assoc
  1012 
  1013 lemma disj_comms:
  1014   shows disj_commute: "(P|Q) = (Q|P)"
  1015     and disj_left_commute: "(P|(Q|R)) = (Q|(P|R))" by iprover+
  1016 lemma disj_assoc: "((P|Q)|R) = (P|(Q|R))" by iprover
  1017 
  1018 lemmas disj_ac = disj_commute disj_left_commute disj_assoc
  1019 
  1020 lemma conj_disj_distribL: "(P&(Q|R)) = (P&Q | P&R)" by iprover
  1021 lemma conj_disj_distribR: "((P|Q)&R) = (P&R | Q&R)" by iprover
  1022 
  1023 lemma disj_conj_distribL: "(P|(Q&R)) = ((P|Q) & (P|R))" by iprover
  1024 lemma disj_conj_distribR: "((P&Q)|R) = ((P|R) & (Q|R))" by iprover
  1025 
  1026 lemma imp_conjR: "(P --> (Q&R)) = ((P-->Q) & (P-->R))" by iprover
  1027 lemma imp_conjL: "((P&Q) -->R)  = (P --> (Q --> R))" by iprover
  1028 lemma imp_disjL: "((P|Q) --> R) = ((P-->R)&(Q-->R))" by iprover
  1029 
  1030 text {* These two are specialized, but @{text imp_disj_not1} is useful in @{text "Auth/Yahalom"}. *}
  1031 lemma imp_disj_not1: "(P --> Q | R) = (~Q --> P --> R)" by blast
  1032 lemma imp_disj_not2: "(P --> Q | R) = (~R --> P --> Q)" by blast
  1033 
  1034 lemma imp_disj1: "((P-->Q)|R) = (P--> Q|R)" by blast
  1035 lemma imp_disj2: "(Q|(P-->R)) = (P--> Q|R)" by blast
  1036 
  1037 lemma imp_cong: "(P = P') ==> (P' ==> (Q = Q')) ==> ((P --> Q) = (P' --> Q'))"
  1038   by iprover
  1039 
  1040 lemma de_Morgan_disj: "(~(P | Q)) = (~P & ~Q)" by iprover
  1041 lemma de_Morgan_conj: "(~(P & Q)) = (~P | ~Q)" by blast
  1042 lemma not_imp: "(~(P --> Q)) = (P & ~Q)" by blast
  1043 lemma not_iff: "(P~=Q) = (P = (~Q))" by blast
  1044 lemma disj_not1: "(~P | Q) = (P --> Q)" by blast
  1045 lemma disj_not2: "(P | ~Q) = (Q --> P)"  -- {* changes orientation :-( *}
  1046   by blast
  1047 lemma imp_conv_disj: "(P --> Q) = ((~P) | Q)" by blast
  1048 
  1049 lemma iff_conv_conj_imp: "(P = Q) = ((P --> Q) & (Q --> P))" by iprover
  1050 
  1051 
  1052 lemma cases_simp: "((P --> Q) & (~P --> Q)) = Q"
  1053   -- {* Avoids duplication of subgoals after @{text split_if}, when the true and false *}
  1054   -- {* cases boil down to the same thing. *}
  1055   by blast
  1056 
  1057 lemma not_all: "(~ (! x. P(x))) = (? x.~P(x))" by blast
  1058 lemma imp_all: "((! x. P x) --> Q) = (? x. P x --> Q)" by blast
  1059 lemma not_ex: "(~ (? x. P(x))) = (! x.~P(x))" by iprover
  1060 lemma imp_ex: "((? x. P x) --> Q) = (! x. P x --> Q)" by iprover
  1061 lemma all_not_ex: "(ALL x. P x) = (~ (EX x. ~ P x ))" by blast
  1062 
  1063 declare All_def [no_atp]
  1064 
  1065 lemma ex_disj_distrib: "(? x. P(x) | Q(x)) = ((? x. P(x)) | (? x. Q(x)))" by iprover
  1066 lemma all_conj_distrib: "(!x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))" by iprover
  1067 
  1068 text {*
  1069   \medskip The @{text "&"} congruence rule: not included by default!
  1070   May slow rewrite proofs down by as much as 50\% *}
  1071 
  1072 lemma conj_cong:
  1073     "(P = P') ==> (P' ==> (Q = Q')) ==> ((P & Q) = (P' & Q'))"
  1074   by iprover
  1075 
  1076 lemma rev_conj_cong:
  1077     "(Q = Q') ==> (Q' ==> (P = P')) ==> ((P & Q) = (P' & Q'))"
  1078   by iprover
  1079 
  1080 text {* The @{text "|"} congruence rule: not included by default! *}
  1081 
  1082 lemma disj_cong:
  1083     "(P = P') ==> (~P' ==> (Q = Q')) ==> ((P | Q) = (P' | Q'))"
  1084   by blast
  1085 
  1086 
  1087 text {* \medskip if-then-else rules *}
  1088 
  1089 lemma if_True [code]: "(if True then x else y) = x"
  1090   by (unfold If_def) blast
  1091 
  1092 lemma if_False [code]: "(if False then x else y) = y"
  1093   by (unfold If_def) blast
  1094 
  1095 lemma if_P: "P ==> (if P then x else y) = x"
  1096   by (unfold If_def) blast
  1097 
  1098 lemma if_not_P: "~P ==> (if P then x else y) = y"
  1099   by (unfold If_def) blast
  1100 
  1101 lemma split_if: "P (if Q then x else y) = ((Q --> P(x)) & (~Q --> P(y)))"
  1102   apply (rule case_split [of Q])
  1103    apply (simplesubst if_P)
  1104     prefer 3 apply (simplesubst if_not_P, blast+)
  1105   done
  1106 
  1107 lemma split_if_asm: "P (if Q then x else y) = (~((Q & ~P x) | (~Q & ~P y)))"
  1108 by (simplesubst split_if, blast)
  1109 
  1110 lemmas if_splits [no_atp] = split_if split_if_asm
  1111 
  1112 lemma if_cancel: "(if c then x else x) = x"
  1113 by (simplesubst split_if, blast)
  1114 
  1115 lemma if_eq_cancel: "(if x = y then y else x) = x"
  1116 by (simplesubst split_if, blast)
  1117 
  1118 lemma if_bool_eq_conj: "(if P then Q else R) = ((P-->Q) & (~P-->R))"
  1119   -- {* This form is useful for expanding @{text "if"}s on the RIGHT of the @{text "==>"} symbol. *}
  1120   by (rule split_if)
  1121 
  1122 lemma if_bool_eq_disj: "(if P then Q else R) = ((P&Q) | (~P&R))"
  1123   -- {* And this form is useful for expanding @{text "if"}s on the LEFT. *}
  1124   apply (simplesubst split_if, blast)
  1125   done
  1126 
  1127 lemma Eq_TrueI: "P ==> P == True" by (unfold atomize_eq) iprover
  1128 lemma Eq_FalseI: "~P ==> P == False" by (unfold atomize_eq) iprover
  1129 
  1130 text {* \medskip let rules for simproc *}
  1131 
  1132 lemma Let_folded: "f x \<equiv> g x \<Longrightarrow>  Let x f \<equiv> Let x g"
  1133   by (unfold Let_def)
  1134 
  1135 lemma Let_unfold: "f x \<equiv> g \<Longrightarrow>  Let x f \<equiv> g"
  1136   by (unfold Let_def)
  1137 
  1138 text {*
  1139   The following copy of the implication operator is useful for
  1140   fine-tuning congruence rules.  It instructs the simplifier to simplify
  1141   its premise.
  1142 *}
  1143 
  1144 definition simp_implies :: "[prop, prop] => prop"  (infixr "=simp=>" 1) where
  1145   "simp_implies \<equiv> op ==>"
  1146 
  1147 lemma simp_impliesI:
  1148   assumes PQ: "(PROP P \<Longrightarrow> PROP Q)"
  1149   shows "PROP P =simp=> PROP Q"
  1150   apply (unfold simp_implies_def)
  1151   apply (rule PQ)
  1152   apply assumption
  1153   done
  1154 
  1155 lemma simp_impliesE:
  1156   assumes PQ: "PROP P =simp=> PROP Q"
  1157   and P: "PROP P"
  1158   and QR: "PROP Q \<Longrightarrow> PROP R"
  1159   shows "PROP R"
  1160   apply (rule QR)
  1161   apply (rule PQ [unfolded simp_implies_def])
  1162   apply (rule P)
  1163   done
  1164 
  1165 lemma simp_implies_cong:
  1166   assumes PP' :"PROP P == PROP P'"
  1167   and P'QQ': "PROP P' ==> (PROP Q == PROP Q')"
  1168   shows "(PROP P =simp=> PROP Q) == (PROP P' =simp=> PROP Q')"
  1169 proof (unfold simp_implies_def, rule equal_intr_rule)
  1170   assume PQ: "PROP P \<Longrightarrow> PROP Q"
  1171   and P': "PROP P'"
  1172   from PP' [symmetric] and P' have "PROP P"
  1173     by (rule equal_elim_rule1)
  1174   then have "PROP Q" by (rule PQ)
  1175   with P'QQ' [OF P'] show "PROP Q'" by (rule equal_elim_rule1)
  1176 next
  1177   assume P'Q': "PROP P' \<Longrightarrow> PROP Q'"
  1178   and P: "PROP P"
  1179   from PP' and P have P': "PROP P'" by (rule equal_elim_rule1)
  1180   then have "PROP Q'" by (rule P'Q')
  1181   with P'QQ' [OF P', symmetric] show "PROP Q"
  1182     by (rule equal_elim_rule1)
  1183 qed
  1184 
  1185 lemma uncurry:
  1186   assumes "P \<longrightarrow> Q \<longrightarrow> R"
  1187   shows "P \<and> Q \<longrightarrow> R"
  1188   using assms by blast
  1189 
  1190 lemma iff_allI:
  1191   assumes "\<And>x. P x = Q x"
  1192   shows "(\<forall>x. P x) = (\<forall>x. Q x)"
  1193   using assms by blast
  1194 
  1195 lemma iff_exI:
  1196   assumes "\<And>x. P x = Q x"
  1197   shows "(\<exists>x. P x) = (\<exists>x. Q x)"
  1198   using assms by blast
  1199 
  1200 lemma all_comm:
  1201   "(\<forall>x y. P x y) = (\<forall>y x. P x y)"
  1202   by blast
  1203 
  1204 lemma ex_comm:
  1205   "(\<exists>x y. P x y) = (\<exists>y x. P x y)"
  1206   by blast
  1207 
  1208 use "Tools/simpdata.ML"
  1209 ML {* open Simpdata *}
  1210 
  1211 setup {*
  1212   Simplifier.method_setup Splitter.split_modifiers
  1213   #> Simplifier.map_simpset (K Simpdata.simpset_simprocs)
  1214   #> Splitter.setup
  1215   #> clasimp_setup
  1216   #> EqSubst.setup
  1217 *}
  1218 
  1219 text {* Simproc for proving @{text "(y = x) == False"} from premise @{text "~(x = y)"}: *}
  1220 
  1221 simproc_setup neq ("x = y") = {* fn _ =>
  1222 let
  1223   val neq_to_EQ_False = @{thm not_sym} RS @{thm Eq_FalseI};
  1224   fun is_neq eq lhs rhs thm =
  1225     (case Thm.prop_of thm of
  1226       _ $ (Not $ (eq' $ l' $ r')) =>
  1227         Not = HOLogic.Not andalso eq' = eq andalso
  1228         r' aconv lhs andalso l' aconv rhs
  1229     | _ => false);
  1230   fun proc ss ct =
  1231     (case Thm.term_of ct of
  1232       eq $ lhs $ rhs =>
  1233         (case find_first (is_neq eq lhs rhs) (Simplifier.prems_of_ss ss) of
  1234           SOME thm => SOME (thm RS neq_to_EQ_False)
  1235         | NONE => NONE)
  1236      | _ => NONE);
  1237 in proc end;
  1238 *}
  1239 
  1240 simproc_setup let_simp ("Let x f") = {*
  1241 let
  1242   val (f_Let_unfold, x_Let_unfold) =
  1243     let val [(_ $ (f $ x) $ _)] = prems_of @{thm Let_unfold}
  1244     in (cterm_of @{theory} f, cterm_of @{theory} x) end
  1245   val (f_Let_folded, x_Let_folded) =
  1246     let val [(_ $ (f $ x) $ _)] = prems_of @{thm Let_folded}
  1247     in (cterm_of @{theory} f, cterm_of @{theory} x) end;
  1248   val g_Let_folded =
  1249     let val [(_ $ _ $ (g $ _))] = prems_of @{thm Let_folded}
  1250     in cterm_of @{theory} g end;
  1251   fun count_loose (Bound i) k = if i >= k then 1 else 0
  1252     | count_loose (s $ t) k = count_loose s k + count_loose t k
  1253     | count_loose (Abs (_, _, t)) k = count_loose  t (k + 1)
  1254     | count_loose _ _ = 0;
  1255   fun is_trivial_let (Const (@{const_name Let}, _) $ x $ t) =
  1256    case t
  1257     of Abs (_, _, t') => count_loose t' 0 <= 1
  1258      | _ => true;
  1259 in fn _ => fn ss => fn ct => if is_trivial_let (Thm.term_of ct)
  1260   then SOME @{thm Let_def} (*no or one ocurrence of bound variable*)
  1261   else let (*Norbert Schirmer's case*)
  1262     val ctxt = Simplifier.the_context ss;
  1263     val thy = ProofContext.theory_of ctxt;
  1264     val t = Thm.term_of ct;
  1265     val ([t'], ctxt') = Variable.import_terms false [t] ctxt;
  1266   in Option.map (hd o Variable.export ctxt' ctxt o single)
  1267     (case t' of Const (@{const_name Let},_) $ x $ f => (* x and f are already in normal form *)
  1268       if is_Free x orelse is_Bound x orelse is_Const x
  1269       then SOME @{thm Let_def}
  1270       else
  1271         let
  1272           val n = case f of (Abs (x, _, _)) => x | _ => "x";
  1273           val cx = cterm_of thy x;
  1274           val {T = xT, ...} = rep_cterm cx;
  1275           val cf = cterm_of thy f;
  1276           val fx_g = Simplifier.rewrite ss (Thm.capply cf cx);
  1277           val (_ $ _ $ g) = prop_of fx_g;
  1278           val g' = abstract_over (x,g);
  1279         in (if (g aconv g')
  1280              then
  1281                 let
  1282                   val rl =
  1283                     cterm_instantiate [(f_Let_unfold, cf), (x_Let_unfold, cx)] @{thm Let_unfold};
  1284                 in SOME (rl OF [fx_g]) end
  1285              else if Term.betapply (f, x) aconv g then NONE (*avoid identity conversion*)
  1286              else let
  1287                    val abs_g'= Abs (n,xT,g');
  1288                    val g'x = abs_g'$x;
  1289                    val g_g'x = Thm.symmetric (Thm.beta_conversion false (cterm_of thy g'x));
  1290                    val rl = cterm_instantiate
  1291                              [(f_Let_folded, cterm_of thy f), (x_Let_folded, cx),
  1292                               (g_Let_folded, cterm_of thy abs_g')]
  1293                              @{thm Let_folded};
  1294                  in SOME (rl OF [Thm.transitive fx_g g_g'x])
  1295                  end)
  1296         end
  1297     | _ => NONE)
  1298   end
  1299 end *}
  1300 
  1301 lemma True_implies_equals: "(True \<Longrightarrow> PROP P) \<equiv> PROP P"
  1302 proof
  1303   assume "True \<Longrightarrow> PROP P"
  1304   from this [OF TrueI] show "PROP P" .
  1305 next
  1306   assume "PROP P"
  1307   then show "PROP P" .
  1308 qed
  1309 
  1310 lemma ex_simps:
  1311   "!!P Q. (EX x. P x & Q)   = ((EX x. P x) & Q)"
  1312   "!!P Q. (EX x. P & Q x)   = (P & (EX x. Q x))"
  1313   "!!P Q. (EX x. P x | Q)   = ((EX x. P x) | Q)"
  1314   "!!P Q. (EX x. P | Q x)   = (P | (EX x. Q x))"
  1315   "!!P Q. (EX x. P x --> Q) = ((ALL x. P x) --> Q)"
  1316   "!!P Q. (EX x. P --> Q x) = (P --> (EX x. Q x))"
  1317   -- {* Miniscoping: pushing in existential quantifiers. *}
  1318   by (iprover | blast)+
  1319 
  1320 lemma all_simps:
  1321   "!!P Q. (ALL x. P x & Q)   = ((ALL x. P x) & Q)"
  1322   "!!P Q. (ALL x. P & Q x)   = (P & (ALL x. Q x))"
  1323   "!!P Q. (ALL x. P x | Q)   = ((ALL x. P x) | Q)"
  1324   "!!P Q. (ALL x. P | Q x)   = (P | (ALL x. Q x))"
  1325   "!!P Q. (ALL x. P x --> Q) = ((EX x. P x) --> Q)"
  1326   "!!P Q. (ALL x. P --> Q x) = (P --> (ALL x. Q x))"
  1327   -- {* Miniscoping: pushing in universal quantifiers. *}
  1328   by (iprover | blast)+
  1329 
  1330 lemmas [simp] =
  1331   triv_forall_equality (*prunes params*)
  1332   True_implies_equals  (*prune asms `True'*)
  1333   if_True
  1334   if_False
  1335   if_cancel
  1336   if_eq_cancel
  1337   imp_disjL
  1338   (*In general it seems wrong to add distributive laws by default: they
  1339     might cause exponential blow-up.  But imp_disjL has been in for a while
  1340     and cannot be removed without affecting existing proofs.  Moreover,
  1341     rewriting by "(P|Q --> R) = ((P-->R)&(Q-->R))" might be justified on the
  1342     grounds that it allows simplification of R in the two cases.*)
  1343   conj_assoc
  1344   disj_assoc
  1345   de_Morgan_conj
  1346   de_Morgan_disj
  1347   imp_disj1
  1348   imp_disj2
  1349   not_imp
  1350   disj_not1
  1351   not_all
  1352   not_ex
  1353   cases_simp
  1354   the_eq_trivial
  1355   the_sym_eq_trivial
  1356   ex_simps
  1357   all_simps
  1358   simp_thms
  1359 
  1360 lemmas [cong] = imp_cong simp_implies_cong
  1361 lemmas [split] = split_if
  1362 
  1363 ML {* val HOL_ss = @{simpset} *}
  1364 
  1365 text {* Simplifies x assuming c and y assuming ~c *}
  1366 lemma if_cong:
  1367   assumes "b = c"
  1368       and "c \<Longrightarrow> x = u"
  1369       and "\<not> c \<Longrightarrow> y = v"
  1370   shows "(if b then x else y) = (if c then u else v)"
  1371   using assms by simp
  1372 
  1373 text {* Prevents simplification of x and y:
  1374   faster and allows the execution of functional programs. *}
  1375 lemma if_weak_cong [cong]:
  1376   assumes "b = c"
  1377   shows "(if b then x else y) = (if c then x else y)"
  1378   using assms by (rule arg_cong)
  1379 
  1380 text {* Prevents simplification of t: much faster *}
  1381 lemma let_weak_cong:
  1382   assumes "a = b"
  1383   shows "(let x = a in t x) = (let x = b in t x)"
  1384   using assms by (rule arg_cong)
  1385 
  1386 text {* To tidy up the result of a simproc.  Only the RHS will be simplified. *}
  1387 lemma eq_cong2:
  1388   assumes "u = u'"
  1389   shows "(t \<equiv> u) \<equiv> (t \<equiv> u')"
  1390   using assms by simp
  1391 
  1392 lemma if_distrib:
  1393   "f (if c then x else y) = (if c then f x else f y)"
  1394   by simp
  1395 
  1396 
  1397 subsubsection {* Generic cases and induction *}
  1398 
  1399 text {* Rule projections: *}
  1400 
  1401 ML {*
  1402 structure Project_Rule = Project_Rule
  1403 (
  1404   val conjunct1 = @{thm conjunct1}
  1405   val conjunct2 = @{thm conjunct2}
  1406   val mp = @{thm mp}
  1407 )
  1408 *}
  1409 
  1410 definition induct_forall where
  1411   "induct_forall P == \<forall>x. P x"
  1412 
  1413 definition induct_implies where
  1414   "induct_implies A B == A \<longrightarrow> B"
  1415 
  1416 definition induct_equal where
  1417   "induct_equal x y == x = y"
  1418 
  1419 definition induct_conj where
  1420   "induct_conj A B == A \<and> B"
  1421 
  1422 definition induct_true where
  1423   "induct_true == True"
  1424 
  1425 definition induct_false where
  1426   "induct_false == False"
  1427 
  1428 lemma induct_forall_eq: "(!!x. P x) == Trueprop (induct_forall (\<lambda>x. P x))"
  1429   by (unfold atomize_all induct_forall_def)
  1430 
  1431 lemma induct_implies_eq: "(A ==> B) == Trueprop (induct_implies A B)"
  1432   by (unfold atomize_imp induct_implies_def)
  1433 
  1434 lemma induct_equal_eq: "(x == y) == Trueprop (induct_equal x y)"
  1435   by (unfold atomize_eq induct_equal_def)
  1436 
  1437 lemma induct_conj_eq: "(A &&& B) == Trueprop (induct_conj A B)"
  1438   by (unfold atomize_conj induct_conj_def)
  1439 
  1440 lemmas induct_atomize' = induct_forall_eq induct_implies_eq induct_conj_eq
  1441 lemmas induct_atomize = induct_atomize' induct_equal_eq
  1442 lemmas induct_rulify' [symmetric, standard] = induct_atomize'
  1443 lemmas induct_rulify [symmetric, standard] = induct_atomize
  1444 lemmas induct_rulify_fallback =
  1445   induct_forall_def induct_implies_def induct_equal_def induct_conj_def
  1446   induct_true_def induct_false_def
  1447 
  1448 
  1449 lemma induct_forall_conj: "induct_forall (\<lambda>x. induct_conj (A x) (B x)) =
  1450     induct_conj (induct_forall A) (induct_forall B)"
  1451   by (unfold induct_forall_def induct_conj_def) iprover
  1452 
  1453 lemma induct_implies_conj: "induct_implies C (induct_conj A B) =
  1454     induct_conj (induct_implies C A) (induct_implies C B)"
  1455   by (unfold induct_implies_def induct_conj_def) iprover
  1456 
  1457 lemma induct_conj_curry: "(induct_conj A B ==> PROP C) == (A ==> B ==> PROP C)"
  1458 proof
  1459   assume r: "induct_conj A B ==> PROP C" and A B
  1460   show "PROP C" by (rule r) (simp add: induct_conj_def `A` `B`)
  1461 next
  1462   assume r: "A ==> B ==> PROP C" and "induct_conj A B"
  1463   show "PROP C" by (rule r) (simp_all add: `induct_conj A B` [unfolded induct_conj_def])
  1464 qed
  1465 
  1466 lemmas induct_conj = induct_forall_conj induct_implies_conj induct_conj_curry
  1467 
  1468 lemma induct_trueI: "induct_true"
  1469   by (simp add: induct_true_def)
  1470 
  1471 text {* Method setup. *}
  1472 
  1473 ML {*
  1474 structure Induct = Induct
  1475 (
  1476   val cases_default = @{thm case_split}
  1477   val atomize = @{thms induct_atomize}
  1478   val rulify = @{thms induct_rulify'}
  1479   val rulify_fallback = @{thms induct_rulify_fallback}
  1480   val equal_def = @{thm induct_equal_def}
  1481   fun dest_def (Const (@{const_name induct_equal}, _) $ t $ u) = SOME (t, u)
  1482     | dest_def _ = NONE
  1483   val trivial_tac = match_tac @{thms induct_trueI}
  1484 )
  1485 *}
  1486 
  1487 setup {*
  1488   Induct.setup #>
  1489   Context.theory_map (Induct.map_simpset (fn ss => ss
  1490     setmksimps (fn ss => Simpdata.mksimps Simpdata.mksimps_pairs ss #>
  1491       map (Simplifier.rewrite_rule (map Thm.symmetric
  1492         @{thms induct_rulify_fallback})))
  1493     addsimprocs
  1494       [Simplifier.simproc_global @{theory} "swap_induct_false"
  1495          ["induct_false ==> PROP P ==> PROP Q"]
  1496          (fn _ => fn _ =>
  1497             (fn _ $ (P as _ $ @{const induct_false}) $ (_ $ Q $ _) =>
  1498                   if P <> Q then SOME Drule.swap_prems_eq else NONE
  1499               | _ => NONE)),
  1500        Simplifier.simproc_global @{theory} "induct_equal_conj_curry"
  1501          ["induct_conj P Q ==> PROP R"]
  1502          (fn _ => fn _ =>
  1503             (fn _ $ (_ $ P) $ _ =>
  1504                 let
  1505                   fun is_conj (@{const induct_conj} $ P $ Q) =
  1506                         is_conj P andalso is_conj Q
  1507                     | is_conj (Const (@{const_name induct_equal}, _) $ _ $ _) = true
  1508                     | is_conj @{const induct_true} = true
  1509                     | is_conj @{const induct_false} = true
  1510                     | is_conj _ = false
  1511                 in if is_conj P then SOME @{thm induct_conj_curry} else NONE end
  1512               | _ => NONE))]))
  1513 *}
  1514 
  1515 text {* Pre-simplification of induction and cases rules *}
  1516 
  1517 lemma [induct_simp]: "(!!x. induct_equal x t ==> PROP P x) == PROP P t"
  1518   unfolding induct_equal_def
  1519 proof
  1520   assume R: "!!x. x = t ==> PROP P x"
  1521   show "PROP P t" by (rule R [OF refl])
  1522 next
  1523   fix x assume "PROP P t" "x = t"
  1524   then show "PROP P x" by simp
  1525 qed
  1526 
  1527 lemma [induct_simp]: "(!!x. induct_equal t x ==> PROP P x) == PROP P t"
  1528   unfolding induct_equal_def
  1529 proof
  1530   assume R: "!!x. t = x ==> PROP P x"
  1531   show "PROP P t" by (rule R [OF refl])
  1532 next
  1533   fix x assume "PROP P t" "t = x"
  1534   then show "PROP P x" by simp
  1535 qed
  1536 
  1537 lemma [induct_simp]: "(induct_false ==> P) == Trueprop induct_true"
  1538   unfolding induct_false_def induct_true_def
  1539   by (iprover intro: equal_intr_rule)
  1540 
  1541 lemma [induct_simp]: "(induct_true ==> PROP P) == PROP P"
  1542   unfolding induct_true_def
  1543 proof
  1544   assume R: "True \<Longrightarrow> PROP P"
  1545   from TrueI show "PROP P" by (rule R)
  1546 next
  1547   assume "PROP P"
  1548   then show "PROP P" .
  1549 qed
  1550 
  1551 lemma [induct_simp]: "(PROP P ==> induct_true) == Trueprop induct_true"
  1552   unfolding induct_true_def
  1553   by (iprover intro: equal_intr_rule)
  1554 
  1555 lemma [induct_simp]: "(!!x. induct_true) == Trueprop induct_true"
  1556   unfolding induct_true_def
  1557   by (iprover intro: equal_intr_rule)
  1558 
  1559 lemma [induct_simp]: "induct_implies induct_true P == P"
  1560   by (simp add: induct_implies_def induct_true_def)
  1561 
  1562 lemma [induct_simp]: "(x = x) = True" 
  1563   by (rule simp_thms)
  1564 
  1565 hide_const induct_forall induct_implies induct_equal induct_conj induct_true induct_false
  1566 
  1567 use "~~/src/Tools/induct_tacs.ML"
  1568 setup InductTacs.setup
  1569 
  1570 
  1571 subsubsection {* Coherent logic *}
  1572 
  1573 ML {*
  1574 structure Coherent = Coherent
  1575 (
  1576   val atomize_elimL = @{thm atomize_elimL}
  1577   val atomize_exL = @{thm atomize_exL}
  1578   val atomize_conjL = @{thm atomize_conjL}
  1579   val atomize_disjL = @{thm atomize_disjL}
  1580   val operator_names =
  1581     [@{const_name "op |"}, @{const_name "op &"}, @{const_name Ex}]
  1582 );
  1583 *}
  1584 
  1585 setup Coherent.setup
  1586 
  1587 
  1588 subsubsection {* Reorienting equalities *}
  1589 
  1590 ML {*
  1591 signature REORIENT_PROC =
  1592 sig
  1593   val add : (term -> bool) -> theory -> theory
  1594   val proc : morphism -> simpset -> cterm -> thm option
  1595 end;
  1596 
  1597 structure Reorient_Proc : REORIENT_PROC =
  1598 struct
  1599   structure Data = Theory_Data
  1600   (
  1601     type T = ((term -> bool) * stamp) list;
  1602     val empty = [];
  1603     val extend = I;
  1604     fun merge data : T = Library.merge (eq_snd op =) data;
  1605   );
  1606   fun add m = Data.map (cons (m, stamp ()));
  1607   fun matches thy t = exists (fn (m, _) => m t) (Data.get thy);
  1608 
  1609   val meta_reorient = @{thm eq_commute [THEN eq_reflection]};
  1610   fun proc phi ss ct =
  1611     let
  1612       val ctxt = Simplifier.the_context ss;
  1613       val thy = ProofContext.theory_of ctxt;
  1614     in
  1615       case Thm.term_of ct of
  1616         (_ $ t $ u) => if matches thy u then NONE else SOME meta_reorient
  1617       | _ => NONE
  1618     end;
  1619 end;
  1620 *}
  1621 
  1622 
  1623 subsection {* Other simple lemmas and lemma duplicates *}
  1624 
  1625 lemma ex1_eq [iff]: "EX! x. x = t" "EX! x. t = x"
  1626   by blast+
  1627 
  1628 lemma choice_eq: "(ALL x. EX! y. P x y) = (EX! f. ALL x. P x (f x))"
  1629   apply (rule iffI)
  1630   apply (rule_tac a = "%x. THE y. P x y" in ex1I)
  1631   apply (fast dest!: theI')
  1632   apply (fast intro: ext the1_equality [symmetric])
  1633   apply (erule ex1E)
  1634   apply (rule allI)
  1635   apply (rule ex1I)
  1636   apply (erule spec)
  1637   apply (erule_tac x = "%z. if z = x then y else f z" in allE)
  1638   apply (erule impE)
  1639   apply (rule allI)
  1640   apply (case_tac "xa = x")
  1641   apply (drule_tac [3] x = x in fun_cong, simp_all)
  1642   done
  1643 
  1644 lemmas eq_sym_conv = eq_commute
  1645 
  1646 lemma nnf_simps:
  1647   "(\<not>(P \<and> Q)) = (\<not> P \<or> \<not> Q)" "(\<not> (P \<or> Q)) = (\<not> P \<and> \<not>Q)" "(P \<longrightarrow> Q) = (\<not>P \<or> Q)" 
  1648   "(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not>(P = Q)) = ((P \<and> \<not> Q) \<or> (\<not>P \<and> Q))" 
  1649   "(\<not> \<not>(P)) = P"
  1650 by blast+
  1651 
  1652 
  1653 subsection {* Basic ML bindings *}
  1654 
  1655 ML {*
  1656 val FalseE = @{thm FalseE}
  1657 val Let_def = @{thm Let_def}
  1658 val TrueI = @{thm TrueI}
  1659 val allE = @{thm allE}
  1660 val allI = @{thm allI}
  1661 val all_dupE = @{thm all_dupE}
  1662 val arg_cong = @{thm arg_cong}
  1663 val box_equals = @{thm box_equals}
  1664 val ccontr = @{thm ccontr}
  1665 val classical = @{thm classical}
  1666 val conjE = @{thm conjE}
  1667 val conjI = @{thm conjI}
  1668 val conjunct1 = @{thm conjunct1}
  1669 val conjunct2 = @{thm conjunct2}
  1670 val disjCI = @{thm disjCI}
  1671 val disjE = @{thm disjE}
  1672 val disjI1 = @{thm disjI1}
  1673 val disjI2 = @{thm disjI2}
  1674 val eq_reflection = @{thm eq_reflection}
  1675 val ex1E = @{thm ex1E}
  1676 val ex1I = @{thm ex1I}
  1677 val ex1_implies_ex = @{thm ex1_implies_ex}
  1678 val exE = @{thm exE}
  1679 val exI = @{thm exI}
  1680 val excluded_middle = @{thm excluded_middle}
  1681 val ext = @{thm ext}
  1682 val fun_cong = @{thm fun_cong}
  1683 val iffD1 = @{thm iffD1}
  1684 val iffD2 = @{thm iffD2}
  1685 val iffI = @{thm iffI}
  1686 val impE = @{thm impE}
  1687 val impI = @{thm impI}
  1688 val meta_eq_to_obj_eq = @{thm meta_eq_to_obj_eq}
  1689 val mp = @{thm mp}
  1690 val notE = @{thm notE}
  1691 val notI = @{thm notI}
  1692 val not_all = @{thm not_all}
  1693 val not_ex = @{thm not_ex}
  1694 val not_iff = @{thm not_iff}
  1695 val not_not = @{thm not_not}
  1696 val not_sym = @{thm not_sym}
  1697 val refl = @{thm refl}
  1698 val rev_mp = @{thm rev_mp}
  1699 val spec = @{thm spec}
  1700 val ssubst = @{thm ssubst}
  1701 val subst = @{thm subst}
  1702 val sym = @{thm sym}
  1703 val trans = @{thm trans}
  1704 *}
  1705 
  1706 
  1707 subsection {* Code generator setup *}
  1708 
  1709 subsubsection {* SML code generator setup *}
  1710 
  1711 use "Tools/recfun_codegen.ML"
  1712 
  1713 setup {*
  1714   Codegen.setup
  1715   #> RecfunCodegen.setup
  1716   #> Codegen.map_unfold (K HOL_basic_ss)
  1717 *}
  1718 
  1719 types_code
  1720   "bool"  ("bool")
  1721 attach (term_of) {*
  1722 fun term_of_bool b = if b then HOLogic.true_const else HOLogic.false_const;
  1723 *}
  1724 attach (test) {*
  1725 fun gen_bool i =
  1726   let val b = one_of [false, true]
  1727   in (b, fn () => term_of_bool b) end;
  1728 *}
  1729   "prop"  ("bool")
  1730 attach (term_of) {*
  1731 fun term_of_prop b =
  1732   HOLogic.mk_Trueprop (if b then HOLogic.true_const else HOLogic.false_const);
  1733 *}
  1734 
  1735 consts_code
  1736   "Trueprop" ("(_)")
  1737   "True"    ("true")
  1738   "False"   ("false")
  1739   "Not"     ("Bool.not")
  1740   "op |"    ("(_ orelse/ _)")
  1741   "op &"    ("(_ andalso/ _)")
  1742   "If"      ("(if _/ then _/ else _)")
  1743 
  1744 setup {*
  1745 let
  1746 
  1747 fun eq_codegen thy defs dep thyname b t gr =
  1748     (case strip_comb t of
  1749        (Const (@{const_name "op ="}, Type (_, [Type ("fun", _), _])), _) => NONE
  1750      | (Const (@{const_name "op ="}, _), [t, u]) =>
  1751           let
  1752             val (pt, gr') = Codegen.invoke_codegen thy defs dep thyname false t gr;
  1753             val (pu, gr'') = Codegen.invoke_codegen thy defs dep thyname false u gr';
  1754             val (_, gr''') = Codegen.invoke_tycodegen thy defs dep thyname false HOLogic.boolT gr'';
  1755           in
  1756             SOME (Codegen.parens
  1757               (Pretty.block [pt, Codegen.str " =", Pretty.brk 1, pu]), gr''')
  1758           end
  1759      | (t as Const (@{const_name "op ="}, _), ts) => SOME (Codegen.invoke_codegen
  1760          thy defs dep thyname b (Codegen.eta_expand t ts 2) gr)
  1761      | _ => NONE);
  1762 
  1763 in
  1764   Codegen.add_codegen "eq_codegen" eq_codegen
  1765 end
  1766 *}
  1767 
  1768 subsubsection {* Generic code generator preprocessor setup *}
  1769 
  1770 setup {*
  1771   Code_Preproc.map_pre (K HOL_basic_ss)
  1772   #> Code_Preproc.map_post (K HOL_basic_ss)
  1773   #> Code_Simp.map_ss (K HOL_basic_ss)
  1774 *}
  1775 
  1776 subsubsection {* Equality *}
  1777 
  1778 class eq =
  1779   fixes eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
  1780   assumes eq_equals: "eq x y \<longleftrightarrow> x = y"
  1781 begin
  1782 
  1783 lemma eq [code_unfold, code_inline del]: "eq = (op =)"
  1784   by (rule ext eq_equals)+
  1785 
  1786 lemma eq_refl: "eq x x \<longleftrightarrow> True"
  1787   unfolding eq by rule+
  1788 
  1789 lemma equals_eq: "(op =) \<equiv> eq"
  1790   by (rule eq_reflection) (rule ext, rule ext, rule sym, rule eq_equals)
  1791 
  1792 declare equals_eq [symmetric, code_post]
  1793 
  1794 end
  1795 
  1796 declare equals_eq [code]
  1797 
  1798 setup {*
  1799   Code_Preproc.map_pre (fn simpset =>
  1800     simpset addsimprocs [Simplifier.simproc_global_i @{theory} "eq" [@{term "op ="}]
  1801       (fn thy => fn _ => fn Const (_, T) => case strip_type T
  1802         of (Type _ :: _, _) => SOME @{thm equals_eq}
  1803          | _ => NONE)])
  1804 *}
  1805 
  1806 
  1807 subsubsection {* Generic code generator foundation *}
  1808 
  1809 text {* Datatypes *}
  1810 
  1811 code_datatype True False
  1812 
  1813 code_datatype "TYPE('a\<Colon>{})"
  1814 
  1815 code_datatype "prop" Trueprop
  1816 
  1817 text {* Code equations *}
  1818 
  1819 lemma [code]:
  1820   shows "(True \<Longrightarrow> PROP Q) \<equiv> PROP Q" 
  1821     and "(PROP Q \<Longrightarrow> True) \<equiv> Trueprop True"
  1822     and "(P \<Longrightarrow> R) \<equiv> Trueprop (P \<longrightarrow> R)" by (auto intro!: equal_intr_rule)
  1823 
  1824 lemma [code]:
  1825   shows "False \<and> P \<longleftrightarrow> False"
  1826     and "True \<and> P \<longleftrightarrow> P"
  1827     and "P \<and> False \<longleftrightarrow> False"
  1828     and "P \<and> True \<longleftrightarrow> P" by simp_all
  1829 
  1830 lemma [code]:
  1831   shows "False \<or> P \<longleftrightarrow> P"
  1832     and "True \<or> P \<longleftrightarrow> True"
  1833     and "P \<or> False \<longleftrightarrow> P"
  1834     and "P \<or> True \<longleftrightarrow> True" by simp_all
  1835 
  1836 lemma [code]:
  1837   shows "(False \<longrightarrow> P) \<longleftrightarrow> True"
  1838     and "(True \<longrightarrow> P) \<longleftrightarrow> P"
  1839     and "(P \<longrightarrow> False) \<longleftrightarrow> \<not> P"
  1840     and "(P \<longrightarrow> True) \<longleftrightarrow> True" by simp_all
  1841 
  1842 instantiation itself :: (type) eq
  1843 begin
  1844 
  1845 definition eq_itself :: "'a itself \<Rightarrow> 'a itself \<Rightarrow> bool" where
  1846   "eq_itself x y \<longleftrightarrow> x = y"
  1847 
  1848 instance proof
  1849 qed (fact eq_itself_def)
  1850 
  1851 end
  1852 
  1853 lemma eq_itself_code [code]:
  1854   "eq_class.eq TYPE('a) TYPE('a) \<longleftrightarrow> True"
  1855   by (simp add: eq)
  1856 
  1857 text {* Equality *}
  1858 
  1859 declare simp_thms(6) [code nbe]
  1860 
  1861 setup {*
  1862   Sign.add_const_constraint (@{const_name eq}, SOME @{typ "'a\<Colon>type \<Rightarrow> 'a \<Rightarrow> bool"})
  1863 *}
  1864 
  1865 lemma equals_alias_cert: "OFCLASS('a, eq_class) \<equiv> ((op = :: 'a \<Rightarrow> 'a \<Rightarrow> bool) \<equiv> eq)" (is "?ofclass \<equiv> ?eq")
  1866 proof
  1867   assume "PROP ?ofclass"
  1868   show "PROP ?eq"
  1869     by (tactic {* ALLGOALS (rtac (Thm.unconstrainT @{thm equals_eq})) *})
  1870       (fact `PROP ?ofclass`)
  1871 next
  1872   assume "PROP ?eq"
  1873   show "PROP ?ofclass" proof
  1874   qed (simp add: `PROP ?eq`)
  1875 qed
  1876   
  1877 setup {*
  1878   Sign.add_const_constraint (@{const_name eq}, SOME @{typ "'a\<Colon>eq \<Rightarrow> 'a \<Rightarrow> bool"})
  1879 *}
  1880 
  1881 setup {*
  1882   Nbe.add_const_alias @{thm equals_alias_cert}
  1883 *}
  1884 
  1885 hide_const (open) eq
  1886 
  1887 text {* Cases *}
  1888 
  1889 lemma Let_case_cert:
  1890   assumes "CASE \<equiv> (\<lambda>x. Let x f)"
  1891   shows "CASE x \<equiv> f x"
  1892   using assms by simp_all
  1893 
  1894 lemma If_case_cert:
  1895   assumes "CASE \<equiv> (\<lambda>b. If b f g)"
  1896   shows "(CASE True \<equiv> f) &&& (CASE False \<equiv> g)"
  1897   using assms by simp_all
  1898 
  1899 setup {*
  1900   Code.add_case @{thm Let_case_cert}
  1901   #> Code.add_case @{thm If_case_cert}
  1902   #> Code.add_undefined @{const_name undefined}
  1903 *}
  1904 
  1905 code_abort undefined
  1906 
  1907 subsubsection {* Generic code generator target languages *}
  1908 
  1909 text {* type bool *}
  1910 
  1911 code_type bool
  1912   (SML "bool")
  1913   (OCaml "bool")
  1914   (Haskell "Bool")
  1915   (Scala "Boolean")
  1916 
  1917 code_const True and False and Not and "op &" and "op |" and If
  1918   (SML "true" and "false" and "not"
  1919     and infixl 1 "andalso" and infixl 0 "orelse"
  1920     and "!(if (_)/ then (_)/ else (_))")
  1921   (OCaml "true" and "false" and "not"
  1922     and infixl 4 "&&" and infixl 2 "||"
  1923     and "!(if (_)/ then (_)/ else (_))")
  1924   (Haskell "True" and "False" and "not"
  1925     and infixl 3 "&&" and infixl 2 "||"
  1926     and "!(if (_)/ then (_)/ else (_))")
  1927   (Scala "true" and "false" and "'! _"
  1928     and infixl 3 "&&" and infixl 1 "||"
  1929     and "!(if ((_))/ (_)/ else (_))")
  1930 
  1931 code_reserved SML
  1932   bool true false not
  1933 
  1934 code_reserved OCaml
  1935   bool not
  1936 
  1937 code_reserved Scala
  1938   Boolean
  1939 
  1940 text {* using built-in Haskell equality *}
  1941 
  1942 code_class eq
  1943   (Haskell "Eq")
  1944 
  1945 code_const "eq_class.eq"
  1946   (Haskell infixl 4 "==")
  1947 
  1948 code_const "op ="
  1949   (Haskell infixl 4 "==")
  1950 
  1951 text {* undefined *}
  1952 
  1953 code_const undefined
  1954   (SML "!(raise/ Fail/ \"undefined\")")
  1955   (OCaml "failwith/ \"undefined\"")
  1956   (Haskell "error/ \"undefined\"")
  1957   (Scala "!error(\"undefined\")")
  1958 
  1959 subsubsection {* Evaluation and normalization by evaluation *}
  1960 
  1961 text {* Avoid some named infixes in evaluation environment *}
  1962 
  1963 code_reserved Eval oo ooo oooo upto downto orf andf
  1964 
  1965 setup {*
  1966   Value.add_evaluator ("SML", Codegen.eval_term o ProofContext.theory_of)
  1967 *}
  1968 
  1969 ML {*
  1970 structure Eval_Method =
  1971 struct
  1972 
  1973 val eval_ref : (unit -> bool) option Unsynchronized.ref = Unsynchronized.ref NONE;
  1974 
  1975 end;
  1976 *}
  1977 
  1978 oracle eval_oracle = {* fn ct =>
  1979   let
  1980     val thy = Thm.theory_of_cterm ct;
  1981     val t = Thm.term_of ct;
  1982     val dummy = @{cprop True};
  1983   in case try HOLogic.dest_Trueprop t
  1984    of SOME t' => if Code_Eval.eval NONE
  1985          ("Eval_Method.eval_ref", Eval_Method.eval_ref) (K I) thy t' [] 
  1986        then Thm.capply (Thm.capply @{cterm "op \<equiv> \<Colon> prop \<Rightarrow> prop \<Rightarrow> prop"} ct) dummy
  1987        else dummy
  1988     | NONE => dummy
  1989   end
  1990 *}
  1991 
  1992 ML {*
  1993 fun gen_eval_method conv ctxt = SIMPLE_METHOD'
  1994   (CONVERSION (Conv.params_conv (~1) (K (Conv.concl_conv (~1) conv)) ctxt)
  1995     THEN' rtac TrueI)
  1996 *}
  1997 
  1998 method_setup eval = {* Scan.succeed (gen_eval_method eval_oracle) *}
  1999   "solve goal by evaluation"
  2000 
  2001 method_setup evaluation = {* Scan.succeed (gen_eval_method Codegen.evaluation_conv) *}
  2002   "solve goal by evaluation"
  2003 
  2004 method_setup normalization = {*
  2005   Scan.succeed (K (SIMPLE_METHOD' (CONVERSION Nbe.dynamic_eval_conv THEN' (fn k => TRY (rtac TrueI k)))))
  2006 *} "solve goal by normalization"
  2007 
  2008 
  2009 subsection {* Counterexample Search Units *}
  2010 
  2011 subsubsection {* Quickcheck *}
  2012 
  2013 quickcheck_params [size = 5, iterations = 50]
  2014 
  2015 
  2016 subsubsection {* Nitpick setup *}
  2017 
  2018 ML {*
  2019 structure Nitpick_Defs = Named_Thms
  2020 (
  2021   val name = "nitpick_def"
  2022   val description = "alternative definitions of constants as needed by Nitpick"
  2023 )
  2024 structure Nitpick_Simps = Named_Thms
  2025 (
  2026   val name = "nitpick_simp"
  2027   val description = "equational specification of constants as needed by Nitpick"
  2028 )
  2029 structure Nitpick_Psimps = Named_Thms
  2030 (
  2031   val name = "nitpick_psimp"
  2032   val description = "partial equational specification of constants as needed by Nitpick"
  2033 )
  2034 structure Nitpick_Choice_Specs = Named_Thms
  2035 (
  2036   val name = "nitpick_choice_spec"
  2037   val description = "choice specification of constants as needed by Nitpick"
  2038 )
  2039 *}
  2040 
  2041 setup {*
  2042   Nitpick_Defs.setup
  2043   #> Nitpick_Simps.setup
  2044   #> Nitpick_Psimps.setup
  2045   #> Nitpick_Choice_Specs.setup
  2046 *}
  2047 
  2048 
  2049 subsection {* Preprocessing for the predicate compiler *}
  2050 
  2051 ML {*
  2052 structure Predicate_Compile_Alternative_Defs = Named_Thms
  2053 (
  2054   val name = "code_pred_def"
  2055   val description = "alternative definitions of constants for the Predicate Compiler"
  2056 )
  2057 structure Predicate_Compile_Inline_Defs = Named_Thms
  2058 (
  2059   val name = "code_pred_inline"
  2060   val description = "inlining definitions for the Predicate Compiler"
  2061 )
  2062 structure Predicate_Compile_Simps = Named_Thms
  2063 (
  2064   val name = "code_pred_simp"
  2065   val description = "simplification rules for the optimisations in the Predicate Compiler"
  2066 )
  2067 *}
  2068 
  2069 setup {*
  2070   Predicate_Compile_Alternative_Defs.setup
  2071   #> Predicate_Compile_Inline_Defs.setup
  2072   #> Predicate_Compile_Simps.setup
  2073 *}
  2074 
  2075 
  2076 subsection {* Legacy tactics and ML bindings *}
  2077 
  2078 ML {*
  2079 fun strip_tac i = REPEAT (resolve_tac [impI, allI] i);
  2080 
  2081 (* combination of (spec RS spec RS ...(j times) ... spec RS mp) *)
  2082 local
  2083   fun wrong_prem (Const (@{const_name All}, _) $ Abs (_, _, t)) = wrong_prem t
  2084     | wrong_prem (Bound _) = true
  2085     | wrong_prem _ = false;
  2086   val filter_right = filter (not o wrong_prem o HOLogic.dest_Trueprop o hd o Thm.prems_of);
  2087 in
  2088   fun smp i = funpow i (fn m => filter_right ([spec] RL m)) ([mp]);
  2089   fun smp_tac j = EVERY'[dresolve_tac (smp j), atac];
  2090 end;
  2091 
  2092 val all_conj_distrib = thm "all_conj_distrib";
  2093 val all_simps = thms "all_simps";
  2094 val atomize_not = thm "atomize_not";
  2095 val case_split = thm "case_split";
  2096 val cases_simp = thm "cases_simp";
  2097 val choice_eq = thm "choice_eq"
  2098 val cong = thm "cong"
  2099 val conj_comms = thms "conj_comms";
  2100 val conj_cong = thm "conj_cong";
  2101 val de_Morgan_conj = thm "de_Morgan_conj";
  2102 val de_Morgan_disj = thm "de_Morgan_disj";
  2103 val disj_assoc = thm "disj_assoc";
  2104 val disj_comms = thms "disj_comms";
  2105 val disj_cong = thm "disj_cong";
  2106 val eq_ac = thms "eq_ac";
  2107 val eq_cong2 = thm "eq_cong2"
  2108 val Eq_FalseI = thm "Eq_FalseI";
  2109 val Eq_TrueI = thm "Eq_TrueI";
  2110 val Ex1_def = thm "Ex1_def"
  2111 val ex_disj_distrib = thm "ex_disj_distrib";
  2112 val ex_simps = thms "ex_simps";
  2113 val if_cancel = thm "if_cancel";
  2114 val if_eq_cancel = thm "if_eq_cancel";
  2115 val if_False = thm "if_False";
  2116 val iff_conv_conj_imp = thm "iff_conv_conj_imp";
  2117 val iff = thm "iff"
  2118 val if_splits = thms "if_splits";
  2119 val if_True = thm "if_True";
  2120 val if_weak_cong = thm "if_weak_cong"
  2121 val imp_all = thm "imp_all";
  2122 val imp_cong = thm "imp_cong";
  2123 val imp_conjL = thm "imp_conjL";
  2124 val imp_conjR = thm "imp_conjR";
  2125 val imp_conv_disj = thm "imp_conv_disj";
  2126 val simp_implies_def = thm "simp_implies_def";
  2127 val simp_thms = thms "simp_thms";
  2128 val split_if = thm "split_if";
  2129 val the1_equality = thm "the1_equality"
  2130 val theI = thm "theI"
  2131 val theI' = thm "theI'"
  2132 val True_implies_equals = thm "True_implies_equals";
  2133 val nnf_conv = Simplifier.rewrite (HOL_basic_ss addsimps simp_thms @ @{thms "nnf_simps"})
  2134 
  2135 *}
  2136 
  2137 end