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