# HG changeset patch # User hoelzl # Date 1295945145 -3600 # Node ID 621fa31e1dbd2d5e8cd14629dca6fce0de5cc6b3 # Parent 1eef159301dfa389f04526cd8a3926a009d558c5# Parent baf1964bc4681a1641a960aac664f3a9b3622699 merged diff -r baf1964bc468 -r 621fa31e1dbd Admin/CHECKLIST --- a/Admin/CHECKLIST Mon Jan 24 22:29:50 2011 +0100 +++ b/Admin/CHECKLIST Tue Jan 25 09:45:45 2011 +0100 @@ -3,7 +3,7 @@ - test polyml-5.4.0, polyml-5.3.0, polyml-5.2.1, smlnj; -- test Proof General 4.1; +- test Proof General 4.1, 4.0, 3.7.1.1; - test Scala wrapper; @@ -46,3 +46,14 @@ - hdiutil create -srcfolder DIR DMG (Mac OS); + +Final release stage +=================== + +- hgrc: default = /home/isabelle-repository/repos/isabelle-release + + isatest@macbroy28:hg-isabelle/.hg/hgrc + isatest@atbroy102:hg-isabelle/.hg/hgrc + +- makedist: REPOS_NAME="isabelle-release" + diff -r baf1964bc468 -r 621fa31e1dbd Admin/isatest/isatest-makedist --- a/Admin/isatest/isatest-makedist Mon Jan 24 22:29:50 2011 +0100 +++ b/Admin/isatest/isatest-makedist Tue Jan 25 09:45:45 2011 +0100 @@ -100,9 +100,9 @@ $SSH macbroy21 "$MAKEALL $HOME/settings/at-poly-test" # give test some time to copy settings and start sleep 15 -$SSH macbroy22 "$MAKEALL $HOME/settings/at-poly" +$SSH macbroy28 "$MAKEALL $HOME/settings/at-poly" sleep 15 -$SSH macbroy28 "$MAKEALL -l HOL HOL-Library $HOME/settings/at-sml-dev-e" +$SSH macbroy22 "$MAKEALL -l HOL HOL-Library $HOME/settings/at-sml-dev-e" sleep 15 $SSH macbroy24 "$MAKEALL $HOME/settings/at64-poly" sleep 15 diff -r baf1964bc468 -r 621fa31e1dbd Admin/makebin --- a/Admin/makebin Mon Jan 24 22:29:50 2011 +0100 +++ b/Admin/makebin Tue Jan 25 09:45:45 2011 +0100 @@ -100,9 +100,7 @@ if [ -n "$DO_LIBRARY" ]; then ./build -bait else - ./build -b -m HOL-Nominal HOL - ./build -b -m HOLCF HOL - ./build -b ZF + ./build -b HOL fi diff -r baf1964bc468 -r 621fa31e1dbd Admin/makebundle --- a/Admin/makebundle Mon Jan 24 22:29:50 2011 +0100 +++ b/Admin/makebundle Tue Jan 25 09:45:45 2011 +0100 @@ -72,7 +72,7 @@ BUNDLE_ARCHIVE="${ARCHIVE_DIR}/${ISABELLE_NAME}_bundle_${PLATFORM}.tar.gz" echo "$(basename "$BUNDLE_ARCHIVE")" -tar -C "$TMP" -c -z -f "$BUNDLE_ARCHIVE" Isabelle "$ISABELLE_NAME" +tar -C "$TMP" -c -z -f "$BUNDLE_ARCHIVE" "$ISABELLE_NAME" # clean up diff -r baf1964bc468 -r 621fa31e1dbd Admin/makedist --- a/Admin/makedist Mon Jan 24 22:29:50 2011 +0100 +++ b/Admin/makedist Tue Jan 25 09:45:45 2011 +0100 @@ -153,7 +153,6 @@ cp doc/isabelle*.eps lib/logo - if [ -z "$RELEASE" ]; then { echo @@ -165,6 +164,7 @@ echo } >ANNOUNCE else + rm Isabelle Isabelle.exe perl -pi -e "s,val is_official = false,val is_official = true,g" src/Pure/ROOT.ML fi diff -r baf1964bc468 -r 621fa31e1dbd Admin/mira.py --- a/Admin/mira.py Mon Jan 24 22:29:50 2011 +0100 +++ b/Admin/mira.py Tue Jan 25 09:45:45 2011 +0100 @@ -10,10 +10,6 @@ import util -from mira.environment import configuration - -from repositories import * - # build and evaluation tools @@ -213,10 +209,9 @@ except IOError: mutabelle_log = '' - attachments = { 'log': log, 'mutabelle_log': mutabelle_log} - return (return_code == 0 and mutabelle_log != '', extract_isabelle_run_summary(log), - {'timing': extract_isabelle_run_timing(log)}, {'log': log}, None) + {'timing': extract_isabelle_run_timing(log)}, + {'log': log, 'mutabelle_log': mutabelle_log}, None) @configuration(repos = [Isabelle], deps = [(HOL, [0])]) def Mutabelle_Relation(*args): diff -r baf1964bc468 -r 621fa31e1dbd Admin/update-keywords --- a/Admin/update-keywords Mon Jan 24 22:29:50 2011 +0100 +++ b/Admin/update-keywords Tue Jan 25 09:45:45 2011 +0100 @@ -12,7 +12,7 @@ isabelle keywords \ "$LOG/Pure.gz" "$LOG/HOL.gz" "$LOG/HOLCF.gz" "$LOG/HOL-Boogie.gz" \ - "$LOG/HOL-Nominal.gz" "$LOG/HOL-Statespace.gz" + "$LOG/HOL-Nominal.gz" "$LOG/HOL-Statespace.gz" "$LOG/HOL-SPARK.gz" isabelle keywords -k ZF \ "$LOG/Pure.gz" "$LOG/FOL.gz" "$LOG/ZF.gz" diff -r baf1964bc468 -r 621fa31e1dbd CONTRIBUTORS --- a/CONTRIBUTORS Mon Jan 24 22:29:50 2011 +0100 +++ b/CONTRIBUTORS Tue Jan 25 09:45:45 2011 +0100 @@ -3,9 +3,16 @@ who is listed as an author in one of the source files of this Isabelle distribution. +Contributions to this Isabelle version +-------------------------------------- + + Contributions to Isabelle2011 ----------------------------- +* January 2011: Stefan Berghofer, secunet Security Networks AG + HOL-SPARK: an interactive prover back-end for SPARK. + * October 2010: Bogdan Grechuk, University of Edinburgh Extended convex analysis in Multivariate Analysis. @@ -17,20 +24,20 @@ partial orders in HOL. * September 2010: Florian Haftmann, TUM - Refined concepts for evaluation, i.e., normalisation of terms using + Refined concepts for evaluation, i.e., normalization of terms using different techniques. * September 2010: Florian Haftmann, TUM Code generation for Scala. * August 2010: Johannes Hoelzl, Armin Heller, and Robert Himmelmann, TUM - Rewriting the Probability theory. + Improved Probability theory in HOL. * July 2010: Florian Haftmann, TUM Reworking and extension of the Imperative HOL framework. -* July 2010: Alexander Krauss, TUM and Christian Sternagel, University of - Innsbruck +* July 2010: Alexander Krauss, TUM and Christian Sternagel, University + of Innsbruck Ad-hoc overloading. Generic do notation for monads. diff -r baf1964bc468 -r 621fa31e1dbd NEWS --- a/NEWS Mon Jan 24 22:29:50 2011 +0100 +++ b/NEWS Tue Jan 25 09:45:45 2011 +0100 @@ -1,13 +1,31 @@ Isabelle NEWS -- history user-relevant changes ============================================== +New in this Isabelle version +---------------------------- + + + New in Isabelle2011 (January 2011) ---------------------------------- *** General *** +* Experimental Prover IDE based on Isabelle/Scala and jEdit (see +src/Tools/jEdit). A bundled component provides "isabelle jedit" as +executable Isabelle tool. Note that this also serves as IDE for +Isabelle/ML, with useful tooltips and hyperlinks produced from its +static analysis. + * Significantly improved Isabelle/Isar implementation manual. +* System settings: ISABELLE_HOME_USER now includes ISABELLE_IDENTIFIER +(and thus refers to something like $HOME/.isabelle/Isabelle2011), +while the default heap location within that directory lacks that extra +suffix. This isolates multiple Isabelle installations from each +other, avoiding problems with old settings in new versions. +INCOMPATIBILITY, need to copy/upgrade old user settings manually. + * Source files are always encoded as UTF-8, instead of old-fashioned ISO-Latin-1. INCOMPATIBILITY. Isabelle LaTeX documents might require the following package declarations: @@ -23,16 +41,11 @@ consistent view on symbols, while raw explode (or String.explode) merely give a byte-oriented representation. -* System settings: ISABELLE_HOME_USER now includes ISABELLE_IDENTIFIER -(and thus refers to something like $HOME/.isabelle/IsabelleXXXX), -while the default heap location within that directory lacks that extra -suffix. This isolates multiple Isabelle installations from each -other, avoiding problems with old settings in new versions. -INCOMPATIBILITY, need to copy/upgrade old user settings manually. - -* Theory loading: only the master source file is looked-up in the -implicit load path, all other files are addressed relatively to its -directory. Minor INCOMPATIBILITY, subtle change in semantics. +* Theory loader: source files are primarily located via the master +directory of each theory node (where the .thy file itself resides). +The global load path is still partially available as legacy feature. +Minor INCOMPATIBILITY due to subtle change in file lookup: use +explicit paths, relatively to the theory. * Special treatment of ML file names has been discontinued. Historically, optional extensions .ML or .sml were added on demand -- @@ -84,9 +97,14 @@ floating-point notation that coincides with the inner syntax for float_token. -* Theory loader: implicit load path is considered legacy. Use -explicit file specifications instead, relatively to the directory of -the enclosing theory file. +* Support for real valued preferences (with approximative PGIP type): +front-ends need to accept "pgint" values in float notation. +INCOMPATIBILITY. + +* The IsabelleText font now includes Cyrillic, Hebrew, Arabic from +DejaVu Sans. + +* Discontinued support for Poly/ML 5.0 and 5.1 versions. *** Pure *** @@ -101,17 +119,15 @@ * Command 'notepad' replaces former 'example_proof' for experimentation in Isar without any result. INCOMPATIBILITY. -* Support for real valued preferences (with approximative PGIP type). - * Locale interpretation commands 'interpret' and 'sublocale' accept lists of equations to map definitions in a locale to appropriate entities in the context of the interpretation. The 'interpretation' command already provided this functionality. -* New diagnostic command 'print_dependencies' prints the locale -instances that would be activated if the specified expression was -interpreted in the current context. Variant 'print_dependencies!' -assumes a context without interpretations. +* Diagnostic command 'print_dependencies' prints the locale instances +that would be activated if the specified expression was interpreted in +the current context. Variant "print_dependencies!" assumes a context +without interpretations. * Diagnostic command 'print_interps' prints interpretations in proofs in addition to interpretations in theories. @@ -124,6 +140,11 @@ * Discontinued obsolete 'constdefs' command. INCOMPATIBILITY, use 'definition' instead. +* The "prems" fact, which refers to the accidental collection of +foundational premises in the context, is now explicitly marked as +legacy feature and will be discontinued soon. Consider using "assms" +of the head statement or reference facts by explicit names. + * Document antiquotations @{class} and @{type} print classes and type constructors. @@ -133,20 +154,14 @@ *** HOL *** -* New simproc that rewrites list comprehensions applied to List.set -to set comprehensions. -To deactivate the simproc for historic proof scripts, use: - - [[simproc del: list_to_set_comprehension]] - -* Functions can be declared as coercions and type inference will add -them as necessary upon input of a term. In theory Complex_Main, -real :: nat => real and real :: int => real are declared as -coercions. A new coercion function f is declared like this: +* Coercive subtyping: functions can be declared as coercions and type +inference will add them as necessary upon input of a term. Theory +Complex_Main declares real :: nat => real and real :: int => real as +coercions. A coercion function f is declared like this: declare [[coercion f]] -To lift coercions through type constructors (eg from nat => real to +To lift coercions through type constructors (e.g. from nat => real to nat list => real list), map functions can be declared, e.g. declare [[coercion_map map]] @@ -158,37 +173,32 @@ declare [[coercion_enabled]] -* New command 'partial_function' provides basic support for recursive -function definitions over complete partial orders. Concrete instances +* Command 'partial_function' provides basic support for recursive +function definitions over complete partial orders. Concrete instances are provided for i) the option type, ii) tail recursion on arbitrary -types, and iii) the heap monad of Imperative_HOL. See -HOL/ex/Fundefs.thy and HOL/Imperative_HOL/ex/Linked_Lists.thy for -examples. - -* Scala (2.8 or higher) has been added to the target languages of the -code generator. - -* Inductive package: offers new command 'inductive_simps' to +types, and iii) the heap monad of Imperative_HOL. See +src/HOL/ex/Fundefs.thy and src/HOL/Imperative_HOL/ex/Linked_Lists.thy +for examples. + +* Function package: f.psimps rules are no longer implicitly declared +as [simp]. INCOMPATIBILITY. + +* Datatype package: theorems generated for executable equality (class +"eq") carry proper names and are treated as default code equations. + +* Inductive package: now offers command 'inductive_simps' to automatically derive instantiated and simplified equations for inductive predicates, similar to 'inductive_cases'. -* Function package: .psimps rules are no longer implicitly declared -[simp]. INCOMPATIBILITY. - -* Datatype package: theorems generated for executable equality (class -eq) carry proper names and are treated as default code equations. - -* New command 'enriched_type' allows to register properties of -the functorial structure of types. - -* Weaker versions of the "meson" and "metis" proof methods are now -available in "HOL-Plain", without dependency on "Hilbert_Choice". The -proof methods become more powerful after "Hilbert_Choice" is loaded in -"HOL-Main". +* Command 'enriched_type' allows to register properties of the +functorial structure of types. * Improved infrastructure for term evaluation using code generator techniques, in particular static evaluation conversions. +* Code generator: Scala (2.8 or higher) has been added to the target +languages. + * Code generator: globbing constant expressions "*" and "Theory.*" have been replaced by the more idiomatic "_" and "Theory._". INCOMPATIBILITY. @@ -199,13 +209,10 @@ * Code generator: do not print function definitions for case combinators any longer. -* Simplification with rules determined by code generator -with code_simp.ML and method code_simp. - -* Records: logical foundation type for records does not carry a '_type' -suffix any longer. INCOMPATIBILITY. - -* Code generation for records: more idiomatic representation of record +* Code generator: simplification with rules determined with +src/Tools/Code/code_simp.ML and method "code_simp". + +* Code generator for records: more idiomatic representation of record types. Warning: records are not covered by ancient SML code generation any longer. INCOMPATIBILITY. In cases of need, a suitable rep_datatype declaration helps to succeed then: @@ -214,16 +221,20 @@ ... rep_datatype foo_ext ... +* Records: logical foundation type for records does not carry a +'_type' suffix any longer (obsolete due to authentic syntax). +INCOMPATIBILITY. + * Quickcheck now by default uses exhaustive testing instead of random -testing. Random testing can be invoked by quickcheck[random], -exhaustive testing by quickcheck[exhaustive]. +testing. Random testing can be invoked by "quickcheck [random]", +exhaustive testing by "quickcheck [exhaustive]". * Quickcheck instantiates polymorphic types with small finite datatypes by default. This enables a simple execution mechanism to handle quantifiers and function equality over the finite datatypes. -* Quickcheck's generator for random generation is renamed from "code" -to "random". INCOMPATIBILITY. +* Quickcheck random generator has been renamed from "code" to +"random". INCOMPATIBILITY. * Quickcheck now has a configurable time limit which is set to 30 seconds by default. This can be changed by adding [timeout = n] to the @@ -234,20 +245,9 @@ counter example search. * Sledgehammer: - - Added "smt" and "remote_smt" provers based on the "smt" proof method. See - the Sledgehammer manual for details ("isabelle doc sledgehammer"). - - Renamed lemmas: - COMBI_def ~> Meson.COMBI_def - COMBK_def ~> Meson.COMBK_def - COMBB_def ~> Meson.COMBB_def - COMBC_def ~> Meson.COMBC_def - COMBS_def ~> Meson.COMBS_def - abs_I ~> Meson.abs_I - abs_K ~> Meson.abs_K - abs_B ~> Meson.abs_B - abs_C ~> Meson.abs_C - abs_S ~> Meson.abs_S - INCOMPATIBILITY. + - Added "smt" and "remote_smt" provers based on the "smt" proof + method. See the Sledgehammer manual for details ("isabelle doc + sledgehammer"). - Renamed commands: sledgehammer atp_info ~> sledgehammer running_provers sledgehammer atp_kill ~> sledgehammer kill_provers @@ -260,18 +260,11 @@ (and "ms" and "min" are no longer supported) INCOMPATIBILITY. -* Metis and Meson now have configuration options "meson_trace", -"metis_trace", and "metis_verbose" that can be enabled to diagnose -these tools. E.g. - - using [[metis_trace = true]] - * Nitpick: - Renamed options: nitpick [timeout = 77 s] ~> nitpick [timeout = 77] nitpick [tac_timeout = 777 ms] ~> nitpick [tac_timeout = 0.777] INCOMPATIBILITY. - - Now requires Kodkodi 1.2.9. INCOMPATIBILITY. - Added support for partial quotient types. - Added local versions of the "Nitpick.register_xxx" functions. - Added "whack" option. @@ -282,19 +275,29 @@ higher cardinalities. - Prevent the expansion of too large definitions. +* Proof methods "metis" and "meson" now have configuration options +"meson_trace", "metis_trace", and "metis_verbose" that can be enabled +to diagnose these tools. E.g. + + using [[metis_trace = true]] + * Auto Solve: Renamed "Auto Solve Direct". The tool is now available manually as command 'solve_direct'. -* The default SMT solver is now CVC3. Z3 must be enabled explicitly, -due to licensing issues. +* The default SMT solver Z3 must be enabled explicitly (due to +licensing issues) by setting the environment variable +Z3_NON_COMMERCIAL in etc/settings of the component, for example. For +commercial applications, the SMT solver CVC3 is provided as fall-back; +changing the SMT solver is done via the configuration option +"smt_solver". * Remote SMT solvers need to be referred to by the "remote_" prefix, -i.e., "remote_cvc3" and "remote_z3". - -* Added basic SMT support for datatypes, records, and typedefs -using the oracle mode (no proofs). Direct support of pairs has been -dropped in exchange (pass theorems fst_conv snd_conv pair_collapse to -the SMT support for a similar behaviour). MINOR INCOMPATIBILITY. +i.e. "remote_cvc3" and "remote_z3". + +* Added basic SMT support for datatypes, records, and typedefs using +the oracle mode (no proofs). Direct support of pairs has been dropped +in exchange (pass theorems fst_conv snd_conv pair_collapse to the SMT +support for a similar behavior). Minor INCOMPATIBILITY. * Changed SMT configuration options: - Renamed: @@ -315,55 +318,13 @@ * Boogie output files (.b2i files) need to be declared in the theory header. -* Dropped syntax for old primrec package. INCOMPATIBILITY. - -* Multivariate Analysis: Introduced a type class for euclidean -space. Most theorems are now stated in terms of euclidean spaces -instead of finite cartesian products. - - types - real ^ 'n ~> 'a::real_vector - ~> 'a::euclidean_space - ~> 'a::ordered_euclidean_space - (depends on your needs) - - constants - _ $ _ ~> _ $$ _ - \ x. _ ~> \\ x. _ - CARD('n) ~> DIM('a) - -Also note that the indices are now natural numbers and not from some -finite type. Finite cartesian products of euclidean spaces, products -of euclidean spaces the real and complex numbers are instantiated to -be euclidean_spaces. INCOMPATIBILITY. - -* Probability: Introduced pextreal as positive extended real numbers. -Use pextreal as value for measures. Introduce the Radon-Nikodym -derivative, product spaces and Fubini's theorem for arbitrary sigma -finite measures. Introduces Lebesgue measure based on the integral in -Multivariate Analysis. INCOMPATIBILITY. - -* Session Imperative_HOL: revamped, corrected dozens of inadequacies. -INCOMPATIBILITY. - -* Theory Library/Monad_Syntax provides do-syntax for monad types. -Syntax in Library/State_Monad has been changed to avoid ambiguities. -INCOMPATIBILITY. - -* Theory SetsAndFunctions has been split into Function_Algebras and -Set_Algebras; canonical names for instance definitions for functions; -various improvements. INCOMPATIBILITY. - -* Theory Multiset provides stable quicksort implementation of -sort_key. - -* Theory Enum (for explicit enumerations of finite types) is now part -of the HOL-Main image. INCOMPATIBILITY: all constants of the Enum -theory now have to be referred to by its qualified name. - - enum ~> Enum.enum - nlists ~> Enum.nlists - product ~> Enum.product +* Simplification procedure "list_to_set_comprehension" rewrites list +comprehensions applied to List.set to set comprehensions. Occasional +INCOMPATIBILITY, may be deactivated like this: + + declare [[simproc del: list_to_set_comprehension]] + +* Removed old version of primrec package. INCOMPATIBILITY. * Removed simplifier congruence rule of "prod_case", as has for long been the case with "split". INCOMPATIBILITY. @@ -373,9 +334,8 @@ * Removed [split_format ... and ... and ...] version of [split_format]. Potential INCOMPATIBILITY. -* Predicate "sorted" now defined inductively, with -nice induction rules. INCOMPATIBILITY: former sorted.simps now -named sorted_simps. +* Predicate "sorted" now defined inductively, with nice induction +rules. INCOMPATIBILITY: former sorted.simps now named sorted_simps. * Constant "contents" renamed to "the_elem", to free the generic name contents for other uses. INCOMPATIBILITY. @@ -386,13 +346,12 @@ * Dropped type classes mult_mono and mult_mono1. INCOMPATIBILITY. -* Removed output syntax "'a ~=> 'b" for "'a => 'b option". INCOMPATIBILITY. +* Removed output syntax "'a ~=> 'b" for "'a => 'b option". +INCOMPATIBILITY. * Renamed theory Fset to Cset, type Fset.fset to Cset.set, in order to avoid confusion with finite sets. INCOMPATIBILITY. -* Multiset.thy: renamed empty_idemp ~> empty_neutral. INCOMPATIBILITY. - * Abandoned locales equiv, congruent and congruent2 for equivalence relations. INCOMPATIBILITY: use equivI rather than equiv_intro (same for congruent(2)). @@ -452,7 +411,7 @@ INCOMPATIBILITY. -* Refactoring of code-generation specific operations in List.thy +* Refactoring of code-generation specific operations in theory List: constants null ~> List.null @@ -468,37 +427,15 @@ Various operations from the Haskell prelude are used for generating Haskell code. -* MESON: Renamed lemmas: - meson_not_conjD ~> Meson.not_conjD - meson_not_disjD ~> Meson.not_disjD - meson_not_notD ~> Meson.not_notD - meson_not_allD ~> Meson.not_allD - meson_not_exD ~> Meson.not_exD - meson_imp_to_disjD ~> Meson.imp_to_disjD - meson_not_impD ~> Meson.not_impD - meson_iff_to_disjD ~> Meson.iff_to_disjD - meson_not_iffD ~> Meson.not_iffD - meson_not_refl_disj_D ~> Meson.not_refl_disj_D - meson_conj_exD1 ~> Meson.conj_exD1 - meson_conj_exD2 ~> Meson.conj_exD2 - meson_disj_exD ~> Meson.disj_exD - meson_disj_exD1 ~> Meson.disj_exD1 - meson_disj_exD2 ~> Meson.disj_exD2 - meson_disj_assoc ~> Meson.disj_assoc - meson_disj_comm ~> Meson.disj_comm - meson_disj_FalseD1 ~> Meson.disj_FalseD1 - meson_disj_FalseD2 ~> Meson.disj_FalseD2 -INCOMPATIBILITY. - -* "bij f" is now an abbreviation of "bij_betw f UNIV UNIV". "surj f" -is now an abbreviation of "range f = UNIV". The theorems bij_def and -surj_def are unchanged. INCOMPATIBILITY. +* Term "bij f" is now an abbreviation of "bij_betw f UNIV UNIV". Term +"surj f" is now an abbreviation of "range f = UNIV". The theorems +bij_def and surj_def are unchanged. INCOMPATIBILITY. * Abolished some non-alphabetic type names: "prod" and "sum" replace "*" and "+" respectively. INCOMPATIBILITY. * Name "Plus" of disjoint sum operator "<+>" is now hidden. Write -Sum_Type.Plus. +"Sum_Type.Plus" instead. * Constant "split" has been merged with constant "prod_case"; names of ML functions, facts etc. involving split have been retained so far, @@ -507,9 +444,65 @@ * Dropped old infix syntax "_ mem _" for List.member; use "_ : set _" instead. INCOMPATIBILITY. -* Removed lemma Option.is_none_none (Duplicate of is_none_def). +* Removed lemma "Option.is_none_none" which duplicates "is_none_def". +INCOMPATIBILITY. + +* Former theory Library/Enum is now part of the HOL-Main image. +INCOMPATIBILITY: all constants of the Enum theory now have to be +referred to by its qualified name. + + enum ~> Enum.enum + nlists ~> Enum.nlists + product ~> Enum.product + +* Theory Library/Monad_Syntax provides do-syntax for monad types. +Syntax in Library/State_Monad has been changed to avoid ambiguities. +INCOMPATIBILITY. + +* Theory Library/SetsAndFunctions has been split into +Library/Function_Algebras and Library/Set_Algebras; canonical names +for instance definitions for functions; various improvements. +INCOMPATIBILITY. + +* Theory Library/Multiset provides stable quicksort implementation of +sort_key. + +* Theory Library/Multiset: renamed empty_idemp ~> empty_neutral. INCOMPATIBILITY. +* Session Multivariate_Analysis: introduced a type class for euclidean +space. Most theorems are now stated in terms of euclidean spaces +instead of finite cartesian products. + + types + real ^ 'n ~> 'a::real_vector + ~> 'a::euclidean_space + ~> 'a::ordered_euclidean_space + (depends on your needs) + + constants + _ $ _ ~> _ $$ _ + \ x. _ ~> \\ x. _ + CARD('n) ~> DIM('a) + +Also note that the indices are now natural numbers and not from some +finite type. Finite cartesian products of euclidean spaces, products +of euclidean spaces the real and complex numbers are instantiated to +be euclidean_spaces. INCOMPATIBILITY. + +* Session Probability: introduced pextreal as positive extended real +numbers. Use pextreal as value for measures. Introduce the +Radon-Nikodym derivative, product spaces and Fubini's theorem for +arbitrary sigma finite measures. Introduces Lebesgue measure based on +the integral in Multivariate Analysis. INCOMPATIBILITY. + +* Session Imperative_HOL: revamped, corrected dozens of inadequacies. +INCOMPATIBILITY. + +* Session SPARK (with image HOL-SPARK) provides commands to load and +prove verification conditions generated by the SPARK Ada program +verifier. See also src/HOL/SPARK and src/HOL/SPARK/Examples. + *** HOL-Algebra *** @@ -519,79 +512,80 @@ qualifier 'add'. Previous theorem names are redeclared for compatibility. -* Structure 'int_ring' is now an abbreviation (previously a +* Structure "int_ring" is now an abbreviation (previously a definition). This fits more natural with advanced interpretations. *** HOLCF *** * The domain package now runs in definitional mode by default: The -former command 'new_domain' is now called 'domain'. To use the domain +former command 'new_domain' is now called 'domain'. To use the domain package in its original axiomatic mode, use 'domain (unsafe)'. INCOMPATIBILITY. -* The new class 'domain' is now the default sort. Class 'predomain' is -an unpointed version of 'domain'. Theories can be updated by replacing -sort annotations as shown below. INCOMPATIBILITY. +* The new class "domain" is now the default sort. Class "predomain" +is an unpointed version of "domain". Theories can be updated by +replacing sort annotations as shown below. INCOMPATIBILITY. 'a::type ~> 'a::countable 'a::cpo ~> 'a::predomain 'a::pcpo ~> 'a::domain -* The old type class 'rep' has been superseded by class 'domain'. +* The old type class "rep" has been superseded by class "domain". Accordingly, users of the definitional package must remove any -'default_sort rep' declarations. INCOMPATIBILITY. +"default_sort rep" declarations. INCOMPATIBILITY. * The domain package (definitional mode) now supports unpointed predomain argument types, as long as they are marked 'lazy'. (Strict -arguments must be in class 'domain'.) For example, the following +arguments must be in class "domain".) For example, the following domain definition now works: domain natlist = nil | cons (lazy "nat discr") (lazy "natlist") * Theory HOLCF/Library/HOL_Cpo provides cpo and predomain class -instances for types from Isabelle/HOL: bool, nat, int, char, 'a + 'b, -'a option, and 'a list. Additionally, it configures fixrec and the -domain package to work with these types. For example: +instances for types from main HOL: bool, nat, int, char, 'a + 'b, +'a option, and 'a list. Additionally, it configures fixrec and the +domain package to work with these types. For example: fixrec isInl :: "('a + 'b) u -> tr" where "isInl$(up$(Inl x)) = TT" | "isInl$(up$(Inr y)) = FF" domain V = VFun (lazy "V -> V") | VCon (lazy "nat") (lazy "V list") -* The '(permissive)' option of fixrec has been replaced with a -per-equation '(unchecked)' option. See HOLCF/Tutorial/Fixrec_ex.thy -for examples. INCOMPATIBILITY. - -* The 'bifinite' class no longer fixes a constant 'approx'; the class -now just asserts that such a function exists. INCOMPATIBILITY. - -* The type 'alg_defl' has been renamed to 'defl'. HOLCF no longer +* The "(permissive)" option of fixrec has been replaced with a +per-equation "(unchecked)" option. See +src/HOL/HOLCF/Tutorial/Fixrec_ex.thy for examples. INCOMPATIBILITY. + +* The "bifinite" class no longer fixes a constant "approx"; the class +now just asserts that such a function exists. INCOMPATIBILITY. + +* Former type "alg_defl" has been renamed to "defl". HOLCF no longer defines an embedding of type 'a defl into udom by default; instances -of 'bifinite' and 'domain' classes are available in -HOLCF/Library/Defl_Bifinite.thy. - -* The syntax 'REP('a)' has been replaced with 'DEFL('a)'. - -* The predicate 'directed' has been removed. INCOMPATIBILITY. - -* The type class 'finite_po' has been removed. INCOMPATIBILITY. - -* The function 'cprod_map' has been renamed to 'prod_map'. +of "bifinite" and "domain" classes are available in +src/HOL/HOLCF/Library/Defl_Bifinite.thy. + +* The syntax "REP('a)" has been replaced with "DEFL('a)". + +* The predicate "directed" has been removed. INCOMPATIBILITY. + +* The type class "finite_po" has been removed. INCOMPATIBILITY. + +* The function "cprod_map" has been renamed to "prod_map". INCOMPATIBILITY. * The monadic bind operator on each powerdomain has new binder syntax -similar to sets, e.g. '\\x\xs. t' represents -'upper_bind\xs\(\ x. t)'. +similar to sets, e.g. "\\x\xs. t" represents +"upper_bind\xs\(\ x. t)". * The infix syntax for binary union on each powerdomain has changed -from e.g. '+\' to '\\', for consistency with set -syntax. INCOMPATIBILITY. - -* The constant 'UU' has been renamed to 'bottom'. The syntax 'UU' is +from e.g. "+\" to "\\", for consistency with set +syntax. INCOMPATIBILITY. + +* The constant "UU" has been renamed to "bottom". The syntax "UU" is still supported as an input translation. * Renamed some theorems (the original names are also still available). + expand_fun_below ~> fun_below_iff below_fun_ext ~> fun_belowI expand_cfun_eq ~> cfun_eq_iff @@ -602,6 +596,7 @@ * The Abs and Rep functions for various types have changed names. Related theorem names have also changed to match. INCOMPATIBILITY. + Rep_CFun ~> Rep_cfun Abs_CFun ~> Abs_cfun Rep_Sprod ~> Rep_sprod @@ -610,20 +605,23 @@ Abs_Ssum ~> Abs_ssum * Lemmas with names of the form *_defined_iff or *_strict_iff have -been renamed to *_bottom_iff. INCOMPATIBILITY. +been renamed to *_bottom_iff. INCOMPATIBILITY. * Various changes to bisimulation/coinduction with domain package: - - Definitions of 'bisim' constants no longer mention definedness. - - With mutual recursion, 'bisim' predicate is now curried. + + - Definitions of "bisim" constants no longer mention definedness. + - With mutual recursion, "bisim" predicate is now curried. - With mutual recursion, each type gets a separate coind theorem. - Variable names in bisim_def and coinduct rules have changed. + INCOMPATIBILITY. -* Case combinators generated by the domain package for type 'foo' are -now named 'foo_case' instead of 'foo_when'. INCOMPATIBILITY. +* Case combinators generated by the domain package for type "foo" are +now named "foo_case" instead of "foo_when". INCOMPATIBILITY. * Several theorems have been renamed to more accurately reflect the -names of constants and types involved. INCOMPATIBILITY. +names of constants and types involved. INCOMPATIBILITY. + thelub_const ~> lub_const lub_const ~> is_lub_const thelubI ~> lub_eqI @@ -645,7 +643,8 @@ deflation_UU ~> deflation_bottom finite_deflation_UU ~> finite_deflation_bottom -* Many legacy theorem names have been discontinued. INCOMPATIBILITY. +* Many legacy theorem names have been discontinued. INCOMPATIBILITY. + sq_ord_less_eq_trans ~> below_eq_trans sq_ord_eq_less_trans ~> eq_below_trans refl_less ~> below_refl @@ -699,11 +698,19 @@ identifiers, e.g. "IFOL.eq" instead of "op =". INCOMPATIBILITY. - *** ML *** -* Renamed structure MetaSimplifier to Raw_Simplifier. Note that the -main functionality is provided by structure Simplifier. +* Antiquotation @{assert} inlines a function bool -> unit that raises +Fail if the argument is false. Due to inlining the source position of +failed assertions is included in the error output. + +* Discontinued antiquotation @{theory_ref}, which is obsolete since ML +text is in practice always evaluated with a stable theory checkpoint. +Minor INCOMPATIBILITY, use (Theory.check_thy @{theory}) instead. + +* Antiquotation @{theory A} refers to theory A from the ancestry of +the current context, not any accidental theory loader state as before. +Potential INCOMPATIBILITY, subtle change in semantics. * Syntax.pretty_priority (default 0) configures the required priority of pretty-printed output and thus affects insertion of parentheses. @@ -714,6 +721,9 @@ * Former exception Library.UnequalLengths now coincides with ListPair.UnequalLengths. +* Renamed structure MetaSimplifier to Raw_Simplifier. Note that the +main functionality is provided by structure Simplifier. + * Renamed raw "explode" function to "raw_explode" to emphasize its meaning. Note that internally to Isabelle, Symbol.explode is used in almost all situations. @@ -722,14 +732,6 @@ See implementation manual for further details on exceptions in Isabelle/ML. -* Antiquotation @{assert} inlines a function bool -> unit that raises -Fail if the argument is false. Due to inlining the source position of -failed assertions is included in the error output. - -* Discontinued antiquotation @{theory_ref}, which is obsolete since ML -text is in practice always evaluated with a stable theory checkpoint. -Minor INCOMPATIBILITY, use (Theory.check_thy @{theory}) instead. - * Renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning. @@ -747,17 +749,12 @@ INCOMPATIBILITY, superseded by static antiquotations @{thm} and @{thms} for most purposes. -* ML structure Unsynchronized never opened, not even in Isar +* ML structure Unsynchronized is never opened, not even in Isar interaction mode as before. Old Unsynchronized.set etc. have been discontinued -- use plain := instead. This should be *rare* anyway, since modern tools always work via official context data, notably configuration options. -* ML antiquotations @{theory} and @{theory_ref} refer to named -theories from the ancestry of the current context, not any accidental -theory loader state as before. Potential INCOMPATIBILITY, subtle -change in semantics. - * Parallel and asynchronous execution requires special care concerning interrupts. Structure Exn provides some convenience functions that avoid working directly with raw Interrupt. User code must not absorb @@ -767,14 +764,6 @@ the program subject to physical effects of the environment. -*** System *** - -* The IsabelleText font now includes Cyrillic, Hebrew, Arabic from -DejaVu Sans. - -* Discontinued support for Poly/ML 5.0 and 5.1 versions. - - New in Isabelle2009-2 (June 2010) --------------------------------- diff -r baf1964bc468 -r 621fa31e1dbd README --- a/README Mon Jan 24 22:29:50 2011 +0100 +++ b/README Tue Jan 25 09:45:45 2011 +0100 @@ -17,6 +17,7 @@ * The GNU bash shell (version 3.x or 2.x). * Perl (version 5.x). * GNU Emacs (version 23) -- for the Proof General 4.x interface. + * Java 1.6.x from Oracle/Sun or Apple -- for Scala and jEdit. * A complete LaTeX installation -- for document preparation. Installation @@ -31,11 +32,15 @@ User interface The classic Isabelle user interface is Proof General by David - Aspinall and others. It is a generic Emacs interface for proof + Aspinall and others. It is a generic Emacs interface for proof assistants, including Isabelle. Its most prominent feature is script management, providing a metaphor of stepwise proof script - editing. Proof General also provides some support for mathematical - symbols displayed on screen. + editing. + + Isabelle/jEdit is an experimental Prover IDE based on advanced + technology of Isabelle/Scala. It provides a metaphor of continuous + proof checking of a versioned collection of theory sources, with + instantaneous feedback in real-time. Other sources of information diff -r baf1964bc468 -r 621fa31e1dbd doc/Contents --- a/doc/Contents Mon Jan 24 22:29:50 2011 +0100 +++ b/doc/Contents Tue Jan 25 09:45:45 2011 +0100 @@ -10,7 +10,7 @@ sledgehammer User's Guide to Sledgehammer sugar LaTeX Sugar for Isabelle documents -Reference Manuals +Main Reference Manuals isar-ref The Isabelle/Isar Reference Manual implementation The Isabelle/Isar Implementation Manual system The Isabelle System Manual diff -r baf1964bc468 -r 621fa31e1dbd etc/isar-keywords.el --- a/etc/isar-keywords.el Mon Jan 24 22:29:50 2011 +0100 +++ b/etc/isar-keywords.el Tue Jan 25 09:45:45 2011 +0100 @@ -1,6 +1,6 @@ ;; ;; Keyword classification tables for Isabelle/Isar. -;; Generated from Pure + HOL + HOLCF + HOL-Boogie + HOL-Nominal + HOL-Statespace. +;; Generated from Pure + HOL + HOLCF + HOL-Boogie + HOL-Nominal + HOL-Statespace + HOL-SPARK. ;; *** DO NOT EDIT *** DO NOT EDIT *** DO NOT EDIT *** ;; @@ -223,6 +223,11 @@ "smt_status" "solve_direct" "sorry" + "spark_end" + "spark_open" + "spark_proof_functions" + "spark_status" + "spark_vc" "specification" "statespace" "subclass" @@ -399,6 +404,7 @@ "sledgehammer" "smt_status" "solve_direct" + "spark_status" "term" "thm" "thm_deps" @@ -511,6 +517,9 @@ "setup" "simproc_setup" "sledgehammer_params" + "spark_end" + "spark_open" + "spark_proof_functions" "statespace" "syntax" "syntax_declaration" @@ -551,6 +560,7 @@ "schematic_corollary" "schematic_lemma" "schematic_theorem" + "spark_vc" "specification" "subclass" "sublocale" diff -r baf1964bc468 -r 621fa31e1dbd etc/proofgeneral-settings.el --- a/etc/proofgeneral-settings.el Mon Jan 24 22:29:50 2011 +0100 +++ b/etc/proofgeneral-settings.el Tue Jan 25 09:45:45 2011 +0100 @@ -3,6 +3,7 @@ ;; Examples for sensible settings: (custom-set-variables '(indent-tabs-mode nil)) +(custom-set-variables '(proof-shell-quit-timeout 45)) ;(custom-set-variables '(isar-eta-contract nil)) diff -r baf1964bc468 -r 621fa31e1dbd lib/html/library_index_content.template --- a/lib/html/library_index_content.template Mon Jan 24 22:29:50 2011 +0100 +++ b/lib/html/library_index_content.template Tue Jan 25 09:45:45 2011 +0100 @@ -7,12 +7,10 @@ is a version of classical higher-order logic resembling that of the HOL System. - - + + diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Big_Operators.thy --- a/src/HOL/Big_Operators.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Big_Operators.thy Tue Jan 25 09:45:45 2011 +0100 @@ -172,7 +172,7 @@ lemma (in comm_monoid_add) setsum_reindex_cong: "[|inj_on f A; B = f ` A; !!a. a:A \ g a = h (f a)|] ==> setsum h B = setsum g A" - by (simp add: setsum_reindex cong: setsum_cong) + by (simp add: setsum_reindex) lemma (in comm_monoid_add) setsum_0[simp]: "setsum (%i. 0) A = 0" by (cases "finite A") (erule finite_induct, auto) @@ -288,7 +288,7 @@ shows "setsum f (UNION I A) = (\i\I. setsum f (A i))" proof - interpret comm_monoid_mult "op +" 0 by (fact comm_monoid_mult) - from assms show ?thesis by (simp add: setsum_def fold_image_UN_disjoint cong: setsum_cong) + from assms show ?thesis by (simp add: setsum_def fold_image_UN_disjoint) qed text{*No need to assume that @{term C} is finite. If infinite, the rhs is @@ -310,7 +310,7 @@ shows "(\x\A. (\y\B x. f x y)) = (\(x,y)\(SIGMA x:A. B x). f x y)" proof - interpret comm_monoid_mult "op +" 0 by (fact comm_monoid_mult) - from assms show ?thesis by (simp add: setsum_def fold_image_Sigma split_def cong: setsum_cong) + from assms show ?thesis by (simp add: setsum_def fold_image_Sigma split_def) qed text{*Here we can eliminate the finiteness assumptions, by cases.*} @@ -498,7 +498,7 @@ assumes "finite A" "A \ {}" and "!!x. x:A \ f x < g x" shows "setsum f A < setsum g A" - using prems + using assms proof (induct rule: finite_ne_induct) case singleton thus ?case by simp next @@ -775,7 +775,7 @@ apply (subgoal_tac "setsum (%i. card (A i)) I = setsum (%i. (setsum (%x. 1) (A i))) I") apply (simp add: setsum_UN_disjoint del: setsum_constant) -apply (simp cong: setsum_cong) +apply simp done lemma card_Union_disjoint: @@ -947,7 +947,7 @@ let ?f = "(\k. if k=a then b k else 1)" {assume a: "a \ S" hence "\ k\ S. ?f k = 1" by simp - hence ?thesis using a by (simp add: setprod_1 cong add: setprod_cong) } + hence ?thesis using a by (simp add: setprod_1) } moreover {assume a: "a \ S" let ?A = "S - {a}" @@ -959,7 +959,7 @@ have "setprod ?f ?A * setprod ?f ?B = setprod ?f S" using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]] by simp - then have ?thesis using a by (simp add: fA1 cong add: setprod_cong cong del: if_weak_cong)} + then have ?thesis using a by (simp add: fA1 cong: setprod_cong cong del: if_weak_cong)} ultimately show ?thesis by blast qed @@ -975,7 +975,7 @@ "finite I ==> (ALL i:I. finite (A i)) ==> (ALL i:I. ALL j:I. i \ j --> A i Int A j = {}) ==> setprod f (UNION I A) = setprod (%i. setprod f (A i)) I" -by(simp add: setprod_def fold_image_UN_disjoint cong: setprod_cong) + by (simp add: setprod_def fold_image_UN_disjoint) lemma setprod_Union_disjoint: "[| (ALL A:C. finite A); @@ -990,7 +990,7 @@ lemma setprod_Sigma: "finite A ==> ALL x:A. finite (B x) ==> (\x\A. (\y\ B x. f x y)) = (\(x,y)\(SIGMA x:A. B x). f x y)" -by(simp add:setprod_def fold_image_Sigma split_def cong:setprod_cong) +by(simp add:setprod_def fold_image_Sigma split_def) text{*Here we can eliminate the finiteness assumptions, by cases.*} lemma setprod_cartesian_product: @@ -1332,7 +1332,7 @@ shows "sup (\\<^bsub>fin\<^esub>A) (\\<^bsub>fin\<^esub>B) = \\<^bsub>fin\<^esub>{sup a b|a b. a \ A \ b \ B}" using A proof (induct rule: finite_ne_induct) case singleton thus ?case - by (simp add: sup_Inf1_distrib [OF B] fold1_singleton_def [OF Inf_fin_def]) + by (simp add: sup_Inf1_distrib [OF B]) next interpret ab_semigroup_idem_mult inf by (rule ab_semigroup_idem_mult_inf) @@ -1347,7 +1347,7 @@ qed have ne: "{sup a b |a b. a \ A \ b \ B} \ {}" using insert B by blast have "sup (\\<^bsub>fin\<^esub>(insert x A)) (\\<^bsub>fin\<^esub>B) = sup (inf x (\\<^bsub>fin\<^esub>A)) (\\<^bsub>fin\<^esub>B)" - using insert by (simp add: fold1_insert_idem_def [OF Inf_fin_def]) + using insert by simp also have "\ = inf (sup x (\\<^bsub>fin\<^esub>B)) (sup (\\<^bsub>fin\<^esub>A) (\\<^bsub>fin\<^esub>B))" by(rule sup_inf_distrib2) also have "\ = inf (\\<^bsub>fin\<^esub>{sup x b|b. b \ B}) (\\<^bsub>fin\<^esub>{sup a b|a b. a \ A \ b \ B})" using insert by(simp add:sup_Inf1_distrib[OF B]) @@ -1391,7 +1391,7 @@ interpret ab_semigroup_idem_mult sup by (rule ab_semigroup_idem_mult_sup) have "inf (\\<^bsub>fin\<^esub>(insert x A)) (\\<^bsub>fin\<^esub>B) = inf (sup x (\\<^bsub>fin\<^esub>A)) (\\<^bsub>fin\<^esub>B)" - using insert by (simp add: fold1_insert_idem_def [OF Sup_fin_def]) + using insert by simp also have "\ = sup (inf x (\\<^bsub>fin\<^esub>B)) (inf (\\<^bsub>fin\<^esub>A) (\\<^bsub>fin\<^esub>B))" by(rule inf_sup_distrib2) also have "\ = sup (\\<^bsub>fin\<^esub>{inf x b|b. b \ B}) (\\<^bsub>fin\<^esub>{inf a b|a b. a \ A \ b \ B})" using insert by(simp add:inf_Sup1_distrib[OF B]) @@ -1551,15 +1551,15 @@ next interpret ab_semigroup_idem_mult min by (rule ab_semigroup_idem_mult_min) - assume "A \ B" + assume neq: "A \ B" have B: "B = A \ (B-A)" using `A \ B` by blast have "fold1 min B = fold1 min (A \ (B-A))" by(subst B)(rule refl) also have "\ = min (fold1 min A) (fold1 min (B-A))" proof - have "finite A" by(rule finite_subset[OF `A \ B` `finite B`]) - moreover have "finite(B-A)" by(rule finite_Diff[OF `finite B`]) (* by(blast intro:finite_Diff prems) fails *) - moreover have "(B-A) \ {}" using prems by blast - moreover have "A Int (B-A) = {}" using prems by blast + moreover have "finite(B-A)" by(rule finite_Diff[OF `finite B`]) + moreover have "(B-A) \ {}" using assms neq by blast + moreover have "A Int (B-A) = {}" using assms by blast ultimately show ?thesis using `A \ {}` by (rule_tac fold1_Un) qed also have "\ \ fold1 min A" by (simp add: min_le_iff_disj) diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Boogie/Examples/Boogie_Dijkstra.thy --- a/src/HOL/Boogie/Examples/Boogie_Dijkstra.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Boogie/Examples/Boogie_Dijkstra.thy Tue Jan 25 09:45:45 2011 +0100 @@ -84,7 +84,7 @@ declare [[smt_certificates="Boogie_Dijkstra.certs"]] declare [[smt_fixed=true]] -declare [[smt_solver=z3, smt_oracle=false]] +declare [[smt_oracle=false]] boogie_vc Dijkstra by boogie diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Boogie/Examples/Boogie_Max.thy --- a/src/HOL/Boogie/Examples/Boogie_Max.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Boogie/Examples/Boogie_Max.thy Tue Jan 25 09:45:45 2011 +0100 @@ -41,7 +41,7 @@ declare [[smt_certificates="Boogie_Max.certs"]] declare [[smt_fixed=true]] -declare [[smt_solver=z3, smt_oracle=false]] +declare [[smt_oracle=false]] boogie_vc max by boogie diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Boogie/Examples/VCC_Max.thy --- a/src/HOL/Boogie/Examples/VCC_Max.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Boogie/Examples/VCC_Max.thy Tue Jan 25 09:45:45 2011 +0100 @@ -49,7 +49,7 @@ declare [[smt_certificates="VCC_Max.certs"]] declare [[smt_fixed=true]] -declare [[smt_solver=z3, smt_oracle=false]] +declare [[smt_oracle=false]] boogie_status diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Deriv.thy Tue Jan 25 09:45:45 2011 +0100 @@ -355,7 +355,7 @@ lemma differentiableE [elim?]: assumes "f differentiable x" obtains df where "DERIV f x :> df" - using prems unfolding differentiable_def .. + using assms unfolding differentiable_def .. lemma differentiableD: "f differentiable x ==> \D. DERIV f x :> D" by (simp add: differentiable_def) @@ -408,7 +408,7 @@ assumes "f differentiable x" assumes "g differentiable x" shows "(\x. f x - g x) differentiable x" - unfolding diff_minus using prems by simp + unfolding diff_minus using assms by simp lemma differentiable_mult [simp]: assumes "f differentiable x" @@ -437,13 +437,16 @@ assumes "f differentiable x" assumes "g differentiable x" and "g x \ 0" shows "(\x. f x / g x) differentiable x" - unfolding divide_inverse using prems by simp + unfolding divide_inverse using assms by simp lemma differentiable_power [simp]: fixes f :: "'a::{real_normed_field} \ 'a" assumes "f differentiable x" shows "(\x. f x ^ n) differentiable x" - by (induct n, simp, simp add: prems) + apply (induct n) + apply simp + apply (simp add: assms) + done subsection {* Nested Intervals and Bisection *} @@ -1227,7 +1230,7 @@ assumes "a < b" and "\x. a \ x & x \ b --> (EX y. DERIV f x :> y & y > 0)" shows "f a < f b" proof (rule ccontr) - assume "~ f a < f b" + assume f: "~ f a < f b" have "EX l z. a < z & z < b & DERIV f z :> l & f b - f a = (b - a) * l" apply (rule MVT) @@ -1236,13 +1239,12 @@ apply (metis DERIV_isCont) apply (metis differentiableI less_le) done - then obtain l z where "a < z" and "z < b" and "DERIV f z :> l" + then obtain l z where z: "a < z" "z < b" "DERIV f z :> l" and "f b - f a = (b - a) * l" by auto - - from prems have "~(l > 0)" + with assms f have "~(l > 0)" by (metis linorder_not_le mult_le_0_iff diff_le_0_iff_le) - with prems show False + with assms z show False by (metis DERIV_unique less_le) qed @@ -1252,9 +1254,8 @@ "\x. a \ x & x \ b --> (\y. DERIV f x :> y & y \ 0)" shows "f a \ f b" proof (rule ccontr, cases "a = b") - assume "~ f a \ f b" - assume "a = b" - with prems show False by auto + assume "~ f a \ f b" and "a = b" + then show False by auto next assume A: "~ f a \ f b" assume B: "a ~= b" @@ -1266,13 +1267,13 @@ apply (metis DERIV_isCont) apply (metis differentiableI less_le) done - then obtain l z where "a < z" and "z < b" and "DERIV f z :> l" + then obtain l z where z: "a < z" "z < b" "DERIV f z :> l" and C: "f b - f a = (b - a) * l" by auto with A have "a < b" "f b < f a" by auto with C have "\ l \ 0" by (auto simp add: not_le algebra_simps) (metis A add_le_cancel_right assms(1) less_eq_real_def mult_right_mono real_add_left_mono real_le_linear real_le_refl) - with prems show False + with assms z show False by (metis DERIV_unique order_less_imp_le) qed @@ -1509,14 +1510,14 @@ theorem GMVT: fixes a b :: real assumes alb: "a < b" - and fc: "\x. a \ x \ x \ b \ isCont f x" - and fd: "\x. a < x \ x < b \ f differentiable x" - and gc: "\x. a \ x \ x \ b \ isCont g x" - and gd: "\x. a < x \ x < b \ g differentiable x" + and fc: "\x. a \ x \ x \ b \ isCont f x" + and fd: "\x. a < x \ x < b \ f differentiable x" + and gc: "\x. a \ x \ x \ b \ isCont g x" + and gd: "\x. a < x \ x < b \ g differentiable x" shows "\g'c f'c c. DERIV g c :> g'c \ DERIV f c :> f'c \ a < c \ c < b \ ((f b - f a) * g'c) = ((g b - g a) * f'c)" proof - let ?h = "\x. (f b - f a)*(g x) - (g b - g a)*(f x)" - from prems have "a < b" by simp + from assms have "a < b" by simp moreover have "\x. a \ x \ x \ b \ isCont ?h x" proof - have "\x. a <= x \ x <= b \ isCont (\x. f b - f a) x" by simp diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Divides.thy --- a/src/HOL/Divides.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Divides.thy Tue Jan 25 09:45:45 2011 +0100 @@ -681,8 +681,8 @@ ML {* local -structure CancelDivMod = CancelDivModFun(struct - +structure CancelDivMod = CancelDivModFun +( val div_name = @{const_name div}; val mod_name = @{const_name mod}; val mk_binop = HOLogic.mk_binop; @@ -691,12 +691,9 @@ val div_mod_eqs = map mk_meta_eq [@{thm div_mod_equality}, @{thm div_mod_equality2}]; - val trans = trans; - val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac (@{thm add_0_left} :: @{thm add_0_right} :: @{thms add_ac})) - -end) +) in @@ -1352,15 +1349,16 @@ theorem posDivAlg_correct: assumes "0 \ a" and "0 < b" shows "divmod_int_rel a b (posDivAlg a b)" -using prems apply (induct a b rule: posDivAlg.induct) -apply auto -apply (simp add: divmod_int_rel_def) -apply (subst posDivAlg_eqn, simp add: right_distrib) -apply (case_tac "a < b") -apply simp_all -apply (erule splitE) -apply (auto simp add: right_distrib Let_def mult_ac mult_2_right) -done + using assms + apply (induct a b rule: posDivAlg.induct) + apply auto + apply (simp add: divmod_int_rel_def) + apply (subst posDivAlg_eqn, simp add: right_distrib) + apply (case_tac "a < b") + apply simp_all + apply (erule splitE) + apply (auto simp add: right_distrib Let_def mult_ac mult_2_right) + done subsubsection{*Correctness of @{term negDivAlg}, the Algorithm for Negative Dividends*} @@ -1381,15 +1379,16 @@ lemma negDivAlg_correct: assumes "a < 0" and "b > 0" shows "divmod_int_rel a b (negDivAlg a b)" -using prems apply (induct a b rule: negDivAlg.induct) -apply (auto simp add: linorder_not_le) -apply (simp add: divmod_int_rel_def) -apply (subst negDivAlg_eqn, assumption) -apply (case_tac "a + b < (0\int)") -apply simp_all -apply (erule splitE) -apply (auto simp add: right_distrib Let_def mult_ac mult_2_right) -done + using assms + apply (induct a b rule: negDivAlg.induct) + apply (auto simp add: linorder_not_le) + apply (simp add: divmod_int_rel_def) + apply (subst negDivAlg_eqn, assumption) + apply (case_tac "a + b < (0\int)") + apply simp_all + apply (erule splitE) + apply (auto simp add: right_distrib Let_def mult_ac mult_2_right) + done subsubsection{*Existence Shown by Proving the Division Algorithm to be Correct*} @@ -1440,8 +1439,8 @@ ML {* local -structure CancelDivMod = CancelDivModFun(struct - +structure CancelDivMod = CancelDivModFun +( val div_name = @{const_name div}; val mod_name = @{const_name mod}; val mk_binop = HOLogic.mk_binop; @@ -1450,12 +1449,9 @@ val div_mod_eqs = map mk_meta_eq [@{thm zdiv_zmod_equality}, @{thm zdiv_zmod_equality2}]; - val trans = trans; - val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac (@{thm diff_minus} :: @{thms add_0s} @ @{thms add_ac})) - -end) +) in diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Fact.thy --- a/src/HOL/Fact.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Fact.thy Tue Jan 25 09:45:45 2011 +0100 @@ -12,12 +12,9 @@ begin class fact = - -fixes - fact :: "'a \ 'a" + fixes fact :: "'a \ 'a" instantiation nat :: fact - begin fun @@ -26,7 +23,7 @@ fact_0_nat: "fact_nat 0 = Suc 0" | fact_Suc: "fact_nat (Suc x) = Suc x * fact x" -instance proof qed +instance .. end @@ -93,8 +90,7 @@ lemma fact_plus_one_int: assumes "n >= 0" shows "fact ((n::int) + 1) = (n + 1) * fact n" - - using prems unfolding fact_int_def + using assms unfolding fact_int_def by (simp add: nat_add_distrib algebra_simps int_mult) lemma fact_reduce_nat: "(n::nat) > 0 \ fact n = n * fact (n - 1)" @@ -102,19 +98,19 @@ apply (erule ssubst) apply (subst fact_Suc) apply simp_all -done + done lemma fact_reduce_int: "(n::int) > 0 \ fact n = n * fact (n - 1)" apply (subgoal_tac "n = (n - 1) + 1") apply (erule ssubst) apply (subst fact_plus_one_int) apply simp_all -done + done lemma fact_nonzero_nat [simp]: "fact (n::nat) \ 0" apply (induct n) apply (auto simp add: fact_plus_one_nat) -done + done lemma fact_nonzero_int [simp]: "n >= 0 \ fact (n::int) ~= 0" by (simp add: fact_int_def) @@ -137,7 +133,7 @@ apply (erule ssubst) apply (subst zle_int) apply auto -done + done lemma dvd_fact_nat [rule_format]: "1 <= m \ m <= n \ m dvd fact (n::nat)" apply (induct n) @@ -147,7 +143,7 @@ apply (erule ssubst) apply (rule dvd_triv_left) apply auto -done + done lemma dvd_fact_int [rule_format]: "1 <= m \ m <= n \ m dvd fact (n::int)" apply (case_tac "1 <= n") @@ -155,7 +151,7 @@ apply (auto simp add: fact_plus_one_int) apply (subgoal_tac "m = i + 1") apply auto -done + done lemma interval_plus_one_nat: "(i::nat) <= j + 1 \ {i..j+1} = {i..j} Un {j+1}" diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Finite_Set.thy Tue Jan 25 09:45:45 2011 +0100 @@ -11,94 +11,48 @@ subsection {* Predicate for finite sets *} -inductive finite :: "'a set => bool" +inductive finite :: "'a set \ bool" where emptyI [simp, intro!]: "finite {}" - | insertI [simp, intro!]: "finite A ==> finite (insert a A)" + | insertI [simp, intro!]: "finite A \ finite (insert a A)" + +lemma finite_induct [case_names empty insert, induct set: finite]: + -- {* Discharging @{text "x \ F"} entails extra work. *} + assumes "finite F" + assumes "P {}" + and insert: "\x F. finite F \ x \ F \ P F \ P (insert x F)" + shows "P F" +using `finite F` proof induct + show "P {}" by fact + fix x F assume F: "finite F" and P: "P F" + show "P (insert x F)" + proof cases + assume "x \ F" + hence "insert x F = F" by (rule insert_absorb) + with P show ?thesis by (simp only:) + next + assume "x \ F" + from F this P show ?thesis by (rule insert) + qed +qed + + +subsubsection {* Choice principles *} lemma ex_new_if_finite: -- "does not depend on def of finite at all" assumes "\ finite (UNIV :: 'a set)" and "finite A" shows "\a::'a. a \ A" proof - from assms have "A \ UNIV" by blast - thus ?thesis by blast -qed - -lemma finite_induct [case_names empty insert, induct set: finite]: - "finite F ==> - P {} ==> (!!x F. finite F ==> x \ F ==> P F ==> P (insert x F)) ==> P F" - -- {* Discharging @{text "x \ F"} entails extra work. *} -proof - - assume "P {}" and - insert: "!!x F. finite F ==> x \ F ==> P F ==> P (insert x F)" - assume "finite F" - thus "P F" - proof induct - show "P {}" by fact - fix x F assume F: "finite F" and P: "P F" - show "P (insert x F)" - proof cases - assume "x \ F" - hence "insert x F = F" by (rule insert_absorb) - with P show ?thesis by (simp only:) - next - assume "x \ F" - from F this P show ?thesis by (rule insert) - qed - qed + then show ?thesis by blast qed -lemma finite_ne_induct[case_names singleton insert, consumes 2]: -assumes fin: "finite F" shows "F \ {} \ - \ \x. P{x}; - \x F. \ finite F; F \ {}; x \ F; P F \ \ P (insert x F) \ - \ P F" -using fin -proof induct - case empty thus ?case by simp -next - case (insert x F) - show ?case - proof cases - assume "F = {}" - thus ?thesis using `P {x}` by simp - next - assume "F \ {}" - thus ?thesis using insert by blast - qed -qed +text {* A finite choice principle. Does not need the SOME choice operator. *} -lemma finite_subset_induct [consumes 2, case_names empty insert]: - assumes "finite F" and "F \ A" - and empty: "P {}" - and insert: "!!a F. finite F ==> a \ A ==> a \ F ==> P F ==> P (insert a F)" - shows "P F" -proof - - from `finite F` and `F \ A` - show ?thesis - proof induct - show "P {}" by fact - next - fix x F - assume "finite F" and "x \ F" and - P: "F \ A ==> P F" and i: "insert x F \ A" - show "P (insert x F)" - proof (rule insert) - from i show "x \ A" by blast - from i have "F \ A" by blast - with P show "P F" . - show "finite F" by fact - show "x \ F" by fact - qed - qed -qed - - -text{* A finite choice principle. Does not need the SOME choice operator. *} lemma finite_set_choice: - "finite A \ ALL x:A. (EX y. P x y) \ EX f. ALL x:A. P x (f x)" -proof (induct set: finite) - case empty thus ?case by simp + "finite A \ \x\A. \y. P x y \ \f. \x\A. P x (f x)" +proof (induct rule: finite_induct) + case empty then show ?case by simp next case (insert a A) then obtain f b where f: "ALL x:A. P x (f x)" and ab: "P a b" by auto @@ -109,16 +63,16 @@ qed -text{* Finite sets are the images of initial segments of natural numbers: *} +subsubsection {* Finite sets are the images of initial segments of natural numbers *} lemma finite_imp_nat_seg_image_inj_on: - assumes fin: "finite A" - shows "\ (n::nat) f. A = f ` {i. i(n::nat) f. A = f ` {i. i < n} \ inj_on f {i. i < n}" +using assms proof induct case empty - show ?case - proof show "\f. {} = f ` {i::nat. i < 0} & inj_on f {i. i<0}" by simp + show ?case + proof + show "\f. {} = f ` {i::nat. i < 0} \ inj_on f {i. i < 0}" by simp qed next case (insert a A) @@ -132,8 +86,8 @@ qed lemma nat_seg_image_imp_finite: - "!!f A. A = f ` {i::nat. i finite A" -proof (induct n) + "A = f ` {i::nat. i < n} \ finite A" +proof (induct n arbitrary: A) case 0 thus ?case by simp next case (Suc n) @@ -152,12 +106,12 @@ qed lemma finite_conv_nat_seg_image: - "finite A = (\ (n::nat) f. A = f ` {i::nat. i (\(n::nat) f. A = f ` {i::nat. i < n})" + by (blast intro: nat_seg_image_imp_finite dest: finite_imp_nat_seg_image_inj_on) lemma finite_imp_inj_to_nat_seg: -assumes "finite A" -shows "EX f n::nat. f`A = {i. if n::nat. f ` A = {i. i < n} \ inj_on f A" proof - from finite_imp_nat_seg_image_inj_on[OF `finite A`] obtain f and n::nat where bij: "bij_betw f {i. i k}" + by (simp add: le_eq_less_or_eq Collect_disj_eq) -lemma finite_UnI: "finite F ==> finite G ==> finite (F Un G)" -by (induct set: finite) simp_all + +subsubsection {* Finiteness and common set operations *} -lemma finite_subset: "A \ B ==> finite B ==> finite A" - -- {* Every subset of a finite set is finite. *} -proof - - assume "finite B" - thus "!!A. A \ B ==> finite A" - proof induct - case empty - thus ?case by simp +lemma rev_finite_subset: + "finite B \ A \ B \ finite A" +proof (induct arbitrary: A rule: finite_induct) + case empty + then show ?case by simp +next + case (insert x F A) + have A: "A \ insert x F" and r: "A - {x} \ F \ finite (A - {x})" by fact+ + show "finite A" + proof cases + assume x: "x \ A" + with A have "A - {x} \ F" by (simp add: subset_insert_iff) + with r have "finite (A - {x})" . + hence "finite (insert x (A - {x}))" .. + also have "insert x (A - {x}) = A" using x by (rule insert_Diff) + finally show ?thesis . next - case (insert x F A) - have A: "A \ insert x F" and r: "A - {x} \ F ==> finite (A - {x})" by fact+ - show "finite A" - proof cases - assume x: "x \ A" - with A have "A - {x} \ F" by (simp add: subset_insert_iff) - with r have "finite (A - {x})" . - hence "finite (insert x (A - {x}))" .. - also have "insert x (A - {x}) = A" using x by (rule insert_Diff) - finally show ?thesis . - next - show "A \ F ==> ?thesis" by fact - assume "x \ A" - with A show "A \ F" by (simp add: subset_insert_iff) - qed + show "A \ F ==> ?thesis" by fact + assume "x \ A" + with A show "A \ F" by (simp add: subset_insert_iff) qed qed -lemma rev_finite_subset: "finite B ==> A \ B ==> finite A" -by (rule finite_subset) - -lemma finite_Un [iff]: "finite (F Un G) = (finite F & finite G)" -by (blast intro: finite_subset [of _ "X Un Y", standard] finite_UnI) - -lemma finite_Collect_disjI[simp]: - "finite{x. P x | Q x} = (finite{x. P x} & finite{x. Q x})" -by(simp add:Collect_disj_eq) - -lemma finite_Int [simp, intro]: "finite F | finite G ==> finite (F Int G)" - -- {* The converse obviously fails. *} -by (blast intro: finite_subset) +lemma finite_subset: + "A \ B \ finite B \ finite A" + by (rule rev_finite_subset) -lemma finite_Collect_conjI [simp, intro]: - "finite{x. P x} | finite{x. Q x} ==> finite{x. P x & Q x}" - -- {* The converse obviously fails. *} -by(simp add:Collect_conj_eq) - -lemma finite_Collect_le_nat[iff]: "finite{n::nat. n<=k}" -by(simp add: le_eq_less_or_eq) - -lemma finite_insert [simp]: "finite (insert a A) = finite A" - apply (subst insert_is_Un) - apply (simp only: finite_Un, blast) - done - -lemma finite_Union[simp, intro]: - "\ finite A; !!M. M \ A \ finite M \ \ finite(\A)" -by (induct rule:finite_induct) simp_all - -lemma finite_Inter[intro]: "EX A:M. finite(A) \ finite(Inter M)" -by (blast intro: Inter_lower finite_subset) +lemma finite_UnI: + assumes "finite F" and "finite G" + shows "finite (F \ G)" + using assms by induct simp_all -lemma finite_INT[intro]: "EX x:I. finite(A x) \ finite(INT x:I. A x)" -by (blast intro: INT_lower finite_subset) +lemma finite_Un [iff]: + "finite (F \ G) \ finite F \ finite G" + by (blast intro: finite_UnI finite_subset [of _ "F \ G"]) -lemma finite_empty_induct: - assumes "finite A" - and "P A" - and "!!a A. finite A ==> a:A ==> P A ==> P (A - {a})" - shows "P {}" +lemma finite_insert [simp]: "finite (insert a A) \ finite A" proof - - have "P (A - A)" - proof - - { - fix c b :: "'a set" - assume c: "finite c" and b: "finite b" - and P1: "P b" and P2: "!!x y. finite y ==> x \ y ==> P y ==> P (y - {x})" - have "c \ b ==> P (b - c)" - using c - proof induct - case empty - from P1 show ?case by simp - next - case (insert x F) - have "P (b - F - {x})" - proof (rule P2) - from _ b show "finite (b - F)" by (rule finite_subset) blast - from insert show "x \ b - F" by simp - from insert show "P (b - F)" by simp - qed - also have "b - F - {x} = b - insert x F" by (rule Diff_insert [symmetric]) - finally show ?case . - qed - } - then show ?thesis by this (simp_all add: assms) - qed + have "finite {a} \ finite A \ finite A" by simp + then have "finite ({a} \ A) \ finite A" by (simp only: finite_Un) then show ?thesis by simp qed -lemma finite_Diff [simp, intro]: "finite A ==> finite (A - B)" -by (rule Diff_subset [THEN finite_subset]) +lemma finite_Int [simp, intro]: + "finite F \ finite G \ finite (F \ G)" + by (blast intro: finite_subset) + +lemma finite_Collect_conjI [simp, intro]: + "finite {x. P x} \ finite {x. Q x} \ finite {x. P x \ Q x}" + by (simp add: Collect_conj_eq) + +lemma finite_Collect_disjI [simp]: + "finite {x. P x \ Q x} \ finite {x. P x} \ finite {x. Q x}" + by (simp add: Collect_disj_eq) + +lemma finite_Diff [simp, intro]: + "finite A \ finite (A - B)" + by (rule finite_subset, rule Diff_subset) lemma finite_Diff2 [simp]: - assumes "finite B" shows "finite (A - B) = finite A" + assumes "finite B" + shows "finite (A - B) \ finite A" proof - - have "finite A \ finite((A-B) Un (A Int B))" by(simp add: Un_Diff_Int) - also have "\ \ finite(A-B)" using `finite B` by(simp) + have "finite A \ finite((A - B) \ (A \ B))" by (simp add: Un_Diff_Int) + also have "\ \ finite (A - B)" using `finite B` by simp finally show ?thesis .. qed +lemma finite_Diff_insert [iff]: + "finite (A - insert a B) \ finite (A - B)" +proof - + have "finite (A - B) \ finite (A - B - {a})" by simp + moreover have "A - insert a B = A - B - {a}" by auto + ultimately show ?thesis by simp +qed + lemma finite_compl[simp]: - "finite(A::'a set) \ finite(-A) = finite(UNIV::'a set)" -by(simp add:Compl_eq_Diff_UNIV) + "finite (A :: 'a set) \ finite (- A) \ finite (UNIV :: 'a set)" + by (simp add: Compl_eq_Diff_UNIV) lemma finite_Collect_not[simp]: - "finite{x::'a. P x} \ finite{x. ~P x} = finite(UNIV::'a set)" -by(simp add:Collect_neg_eq) + "finite {x :: 'a. P x} \ finite {x. \ P x} \ finite (UNIV :: 'a set)" + by (simp add: Collect_neg_eq) + +lemma finite_Union [simp, intro]: + "finite A \ (\M. M \ A \ finite M) \ finite(\A)" + by (induct rule: finite_induct) simp_all + +lemma finite_UN_I [intro]: + "finite A \ (\a. a \ A \ finite (B a)) \ finite (\a\A. B a)" + by (induct rule: finite_induct) simp_all -lemma finite_Diff_insert [iff]: "finite (A - insert a B) = finite (A - B)" - apply (subst Diff_insert) - apply (case_tac "a : A - B") - apply (rule finite_insert [symmetric, THEN trans]) - apply (subst insert_Diff, simp_all) - done +lemma finite_UN [simp]: + "finite A \ finite (UNION A B) \ (\x\A. finite (B x))" + by (blast intro: finite_subset) + +lemma finite_Inter [intro]: + "\A\M. finite A \ finite (\M)" + by (blast intro: Inter_lower finite_subset) - -text {* Image and Inverse Image over Finite Sets *} +lemma finite_INT [intro]: + "\x\I. finite (A x) \ finite (\x\I. A x)" + by (blast intro: INT_lower finite_subset) -lemma finite_imageI[simp, intro]: "finite F ==> finite (h ` F)" - -- {* The image of a finite set is finite. *} - by (induct set: finite) simp_all +lemma finite_imageI [simp, intro]: + "finite F \ finite (h ` F)" + by (induct rule: finite_induct) simp_all lemma finite_image_set [simp]: "finite {x. P x} \ finite { f x | x. P x }" by (simp add: image_Collect [symmetric]) -lemma finite_surj: "finite A ==> B <= f ` A ==> finite B" - apply (frule finite_imageI) - apply (erule finite_subset, assumption) - done - -lemma finite_range_imageI: - "finite (range g) ==> finite (range (%x. f (g x)))" - apply (drule finite_imageI, simp add: range_composition) - done - -lemma finite_imageD: "finite (f`A) ==> inj_on f A ==> finite A" +lemma finite_imageD: + "finite (f ` A) \ inj_on f A \ finite A" proof - have aux: "!!A. finite (A - {}) = finite A" by simp fix B :: "'a set" @@ -340,18 +265,28 @@ done qed (rule refl) +lemma finite_surj: + "finite A \ B \ f ` A \ finite B" + by (erule finite_subset) (rule finite_imageI) -lemma inj_vimage_singleton: "inj f ==> f-`{a} \ {THE x. f x = a}" - -- {* The inverse image of a singleton under an injective function - is included in a singleton. *} - apply (auto simp add: inj_on_def) - apply (blast intro: the_equality [symmetric]) - done +lemma finite_range_imageI: + "finite (range g) \ finite (range (\x. f (g x)))" + by (drule finite_imageI) (simp add: range_composition) -lemma finite_vimageI: "[|finite F; inj h|] ==> finite (h -` F)" - -- {* The inverse image of a finite set under an injective function - is finite. *} - apply (induct set: finite) +lemma finite_subset_image: + assumes "finite B" + shows "B \ f ` A \ \C\A. finite C \ B = f ` C" +using assms proof induct + case empty then show ?case by simp +next + case insert then show ?case + by (clarsimp simp del: image_insert simp add: image_insert [symmetric]) + blast +qed + +lemma finite_vimageI: + "finite F \ inj h \ finite (h -` F)" + apply (induct rule: finite_induct) apply simp_all apply (subst vimage_insert) apply (simp add: finite_subset [OF inj_vimage_singleton]) @@ -369,40 +304,25 @@ lemma finite_vimage_iff: "bij h \ finite (h -` F) \ finite F" unfolding bij_def by (auto elim: finite_vimageD finite_vimageI) - -text {* The finite UNION of finite sets *} - -lemma finite_UN_I[intro]: - "finite A ==> (!!a. a:A ==> finite (B a)) ==> finite (UN a:A. B a)" -by (induct set: finite) simp_all - -text {* - Strengthen RHS to - @{prop "((ALL x:A. finite (B x)) & finite {x. x:A & B x \ {}})"}? - - We'd need to prove - @{prop "finite C ==> ALL A B. (UNION A B) <= C --> finite {x. x:A & B x \ {}}"} - by induction. *} +lemma finite_Collect_bex [simp]: + assumes "finite A" + shows "finite {x. \y\A. Q x y} \ (\y\A. finite {x. Q x y})" +proof - + have "{x. \y\A. Q x y} = (\y\A. {x. Q x y})" by auto + with assms show ?thesis by simp +qed -lemma finite_UN [simp]: - "finite A ==> finite (UNION A B) = (ALL x:A. finite (B x))" -by (blast intro: finite_subset) - -lemma finite_Collect_bex[simp]: "finite A \ - finite{x. EX y:A. Q x y} = (ALL y:A. finite{x. Q x y})" -apply(subgoal_tac "{x. EX y:A. Q x y} = UNION A (%y. {x. Q x y})") - apply auto -done +lemma finite_Collect_bounded_ex [simp]: + assumes "finite {y. P y}" + shows "finite {x. \y. P y \ Q x y} \ (\y. P y \ finite {x. Q x y})" +proof - + have "{x. EX y. P y & Q x y} = (\y\{y. P y}. {x. Q x y})" by auto + with assms show ?thesis by simp +qed -lemma finite_Collect_bounded_ex[simp]: "finite{y. P y} \ - finite{x. EX y. P y & Q x y} = (ALL y. P y \ finite{x. Q x y})" -apply(subgoal_tac "{x. EX y. P y & Q x y} = UNION {y. P y} (%y. {x. Q x y})") - apply auto -done - - -lemma finite_Plus: "[| finite A; finite B |] ==> finite (A <+> B)" -by (simp add: Plus_def) +lemma finite_Plus: + "finite A \ finite B \ finite (A <+> B)" + by (simp add: Plus_def) lemma finite_PlusD: fixes A :: "'a set" and B :: "'b set" @@ -410,42 +330,36 @@ shows "finite A" "finite B" proof - have "Inl ` A \ A <+> B" by auto - hence "finite (Inl ` A :: ('a + 'b) set)" using fin by(rule finite_subset) - thus "finite A" by(rule finite_imageD)(auto intro: inj_onI) + then have "finite (Inl ` A :: ('a + 'b) set)" using fin by (rule finite_subset) + then show "finite A" by (rule finite_imageD) (auto intro: inj_onI) next have "Inr ` B \ A <+> B" by auto - hence "finite (Inr ` B :: ('a + 'b) set)" using fin by(rule finite_subset) - thus "finite B" by(rule finite_imageD)(auto intro: inj_onI) + then have "finite (Inr ` B :: ('a + 'b) set)" using fin by (rule finite_subset) + then show "finite B" by (rule finite_imageD) (auto intro: inj_onI) qed -lemma finite_Plus_iff[simp]: "finite (A <+> B) \ finite A \ finite B" -by(auto intro: finite_PlusD finite_Plus) +lemma finite_Plus_iff [simp]: + "finite (A <+> B) \ finite A \ finite B" + by (auto intro: finite_PlusD finite_Plus) -lemma finite_Plus_UNIV_iff[simp]: - "finite (UNIV :: ('a + 'b) set) = - (finite (UNIV :: 'a set) & finite (UNIV :: 'b set))" -by(subst UNIV_Plus_UNIV[symmetric])(rule finite_Plus_iff) - - -text {* Sigma of finite sets *} +lemma finite_Plus_UNIV_iff [simp]: + "finite (UNIV :: ('a + 'b) set) \ finite (UNIV :: 'a set) \ finite (UNIV :: 'b set)" + by (subst UNIV_Plus_UNIV [symmetric]) (rule finite_Plus_iff) lemma finite_SigmaI [simp, intro]: - "finite A ==> (!!a. a:A ==> finite (B a)) ==> finite (SIGMA a:A. B a)" + "finite A \ (\a. a\A \ finite (B a)) ==> finite (SIGMA a:A. B a)" by (unfold Sigma_def) blast -lemma finite_cartesian_product: "[| finite A; finite B |] ==> - finite (A <*> B)" +lemma finite_cartesian_product: + "finite A \ finite B \ finite (A \ B)" by (rule finite_SigmaI) lemma finite_Prod_UNIV: - "finite (UNIV::'a set) ==> finite (UNIV::'b set) ==> finite (UNIV::('a * 'b) set)" - apply (subgoal_tac "(UNIV:: ('a * 'b) set) = Sigma UNIV (%x. UNIV)") - apply (erule ssubst) - apply (erule finite_SigmaI, auto) - done + "finite (UNIV :: 'a set) \ finite (UNIV :: 'b set) \ finite (UNIV :: ('a \ 'b) set)" + by (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product) lemma finite_cartesian_productD1: - "[| finite (A <*> B); B \ {} |] ==> finite A" + "finite (A \ B) \ B \ {} \ finite A" apply (auto simp add: finite_conv_nat_seg_image) apply (drule_tac x=n in spec) apply (drule_tac x="fst o f" in spec) @@ -474,37 +388,89 @@ apply (rule_tac x=k in image_eqI, auto) done - -text {* The powerset of a finite set *} - -lemma finite_Pow_iff [iff]: "finite (Pow A) = finite A" +lemma finite_Pow_iff [iff]: + "finite (Pow A) \ finite A" proof assume "finite (Pow A)" - with _ have "finite ((%x. {x}) ` A)" by (rule finite_subset) blast - thus "finite A" by (rule finite_imageD [unfolded inj_on_def]) simp + then have "finite ((%x. {x}) ` A)" by (blast intro: finite_subset) + then show "finite A" by (rule finite_imageD [unfolded inj_on_def]) simp next assume "finite A" - thus "finite (Pow A)" + then show "finite (Pow A)" by induct (simp_all add: Pow_insert) qed -lemma finite_Collect_subsets[simp,intro]: "finite A \ finite{B. B \ A}" -by(simp add: Pow_def[symmetric]) - +corollary finite_Collect_subsets [simp, intro]: + "finite A \ finite {B. B \ A}" + by (simp add: Pow_def [symmetric]) lemma finite_UnionD: "finite(\A) \ finite A" -by(blast intro: finite_subset[OF subset_Pow_Union]) + by (blast intro: finite_subset [OF subset_Pow_Union]) -lemma finite_subset_image: - assumes "finite B" - shows "B \ f ` A \ \C\A. finite C \ B = f ` C" -using assms proof(induct) - case empty thus ?case by simp +subsubsection {* Further induction rules on finite sets *} + +lemma finite_ne_induct [case_names singleton insert, consumes 2]: + assumes "finite F" and "F \ {}" + assumes "\x. P {x}" + and "\x F. finite F \ F \ {} \ x \ F \ P F \ P (insert x F)" + shows "P F" +using assms proof induct + case empty then show ?case by simp +next + case (insert x F) then show ?case by cases auto +qed + +lemma finite_subset_induct [consumes 2, case_names empty insert]: + assumes "finite F" and "F \ A" + assumes empty: "P {}" + and insert: "\a F. finite F \ a \ A \ a \ F \ P F \ P (insert a F)" + shows "P F" +using `finite F` `F \ A` proof induct + show "P {}" by fact next - case insert thus ?case - by (clarsimp simp del: image_insert simp add: image_insert[symmetric]) - blast + fix x F + assume "finite F" and "x \ F" and + P: "F \ A \ P F" and i: "insert x F \ A" + show "P (insert x F)" + proof (rule insert) + from i show "x \ A" by blast + from i have "F \ A" by blast + with P show "P F" . + show "finite F" by fact + show "x \ F" by fact + qed +qed + +lemma finite_empty_induct: + assumes "finite A" + assumes "P A" + and remove: "\a A. finite A \ a \ A \ P A \ P (A - {a})" + shows "P {}" +proof - + have "\B. B \ A \ P (A - B)" + proof - + fix B :: "'a set" + assume "B \ A" + with `finite A` have "finite B" by (rule rev_finite_subset) + from this `B \ A` show "P (A - B)" + proof induct + case empty + from `P A` show ?case by simp + next + case (insert b B) + have "P (A - B - {b})" + proof (rule remove) + from `finite A` show "finite (A - B)" by induct auto + from insert show "b \ A - B" by simp + from insert show "P (A - B)" by simp + qed + also have "A - B - {b} = A - insert b B" by (rule Diff_insert [symmetric]) + finally show ?case . + qed + qed + then have "P (A - A)" by blast + then show ?thesis by simp qed @@ -610,7 +576,7 @@ by (induct set: fold_graph) auto lemma finite_imp_fold_graph: "finite A \ \x. fold_graph f z A x" -by (induct set: finite) auto +by (induct rule: finite_induct) auto subsubsection{*From @{const fold_graph} to @{term fold}*} @@ -803,7 +769,7 @@ proof - interpret fun_left_comm_idem inf by (fact fun_left_comm_idem_inf) from `finite A` show ?thesis by (induct A arbitrary: B) - (simp_all add: Inf_empty Inf_insert inf_commute fold_fun_comm) + (simp_all add: Inf_insert inf_commute fold_fun_comm) qed lemma sup_Sup_fold_sup: @@ -812,7 +778,7 @@ proof - interpret fun_left_comm_idem sup by (fact fun_left_comm_idem_sup) from `finite A` show ?thesis by (induct A arbitrary: B) - (simp_all add: Sup_empty Sup_insert sup_commute fold_fun_comm) + (simp_all add: Sup_insert sup_commute fold_fun_comm) qed lemma Inf_fold_inf: @@ -833,7 +799,7 @@ interpret fun_left_comm_idem "\A. inf (f A)" by (fact fun_left_comm_idem_apply) from `finite A` show "?fold = ?inf" by (induct A arbitrary: B) - (simp_all add: INFI_def Inf_empty Inf_insert inf_left_commute) + (simp_all add: INFI_def Inf_insert inf_left_commute) qed lemma sup_SUPR_fold_sup: @@ -844,7 +810,7 @@ interpret fun_left_comm_idem "\A. sup (f A)" by (fact fun_left_comm_idem_apply) from `finite A` show "?fold = ?sup" by (induct A arbitrary: B) - (simp_all add: SUPR_def Sup_empty Sup_insert sup_left_commute) + (simp_all add: SUPR_def Sup_insert sup_left_commute) qed lemma INFI_fold_inf: @@ -949,13 +915,14 @@ lemma fold_image_1: "finite S \ (\x\S. f x = 1) \ fold_image op * f 1 S = 1" - apply (induct set: finite) + apply (induct rule: finite_induct) apply simp by auto lemma fold_image_Un_Int: "finite A ==> finite B ==> fold_image times g 1 A * fold_image times g 1 B = fold_image times g 1 (A Un B) * fold_image times g 1 (A Int B)" + apply (induct rule: finite_induct) by (induct set: finite) (auto simp add: mult_ac insert_absorb Int_insert_left) @@ -981,7 +948,9 @@ ALL i:I. ALL j:I. i \ j --> A i Int A j = {} \ \ fold_image times g 1 (UNION I A) = fold_image times (%i. fold_image times g 1 (A i)) 1 I" -apply (induct set: finite, simp, atomize) +apply (induct rule: finite_induct) +apply simp +apply atomize apply (subgoal_tac "ALL i:F. x \ i") prefer 2 apply blast apply (subgoal_tac "A x Int UNION F A = {}") @@ -1197,19 +1166,19 @@ by (auto simp add: nonempty_iff) show ?thesis proof cases - assume "a = x" - thus ?thesis + assume a: "a = x" + show ?thesis proof cases assume "A' = {}" - with prems show ?thesis by simp + with A' a show ?thesis by simp next assume "A' \ {}" - with prems show ?thesis + with A A' a show ?thesis by (simp add: fold1_insert mult_assoc [symmetric]) qed next assume "a \ x" - with prems show ?thesis + with A A' show ?thesis by (simp add: insert_commute fold1_eq_fold) qed qed @@ -1599,7 +1568,9 @@ and "\i\I. \j\I. i \ j \ A i \ A j = {}" shows "F g (UNION I A) = F (F g \ A) I" apply (insert assms) -apply (induct set: finite, simp, atomize) +apply (induct rule: finite_induct) +apply simp +apply atomize apply (subgoal_tac "\i\Fa. x \ i") prefer 2 apply blast apply (subgoal_tac "A x Int UNION Fa A = {}") @@ -1975,7 +1946,9 @@ qed lemma card_seteq: "finite B ==> (!!A. A <= B ==> card B <= card A ==> A = B)" -apply (induct set: finite, simp, clarify) +apply (induct rule: finite_induct) +apply simp +apply clarify apply (subgoal_tac "finite A & A - {x} <= F") prefer 2 apply (blast intro: finite_subset, atomize) apply (drule_tac x = "A - {x}" in spec) @@ -2146,7 +2119,7 @@ subsubsection {* Cardinality of image *} lemma card_image_le: "finite A ==> card (f ` A) <= card A" -apply (induct set: finite) +apply (induct rule: finite_induct) apply simp apply (simp add: le_SucI card_insert_if) done @@ -2198,6 +2171,7 @@ using assms unfolding bij_betw_def using finite_imageD[of f A] by auto + subsubsection {* Pigeonhole Principles *} lemma pigeonhole: "card A > card(f ` A) \ ~ inj_on f A " @@ -2267,7 +2241,7 @@ subsubsection {* Cardinality of the Powerset *} lemma card_Pow: "finite A ==> card (Pow A) = Suc (Suc 0) ^ card A" (* FIXME numeral 2 (!?) *) -apply (induct set: finite) +apply (induct rule: finite_induct) apply (simp_all add: Pow_insert) apply (subst card_Un_disjoint, blast) apply (blast, blast) @@ -2284,9 +2258,11 @@ ALL c : C. k dvd card c ==> (ALL c1: C. ALL c2: C. c1 \ c2 --> c1 Int c2 = {}) ==> k dvd card (Union C)" -apply(frule finite_UnionD) -apply(rotate_tac -1) -apply (induct set: finite, simp_all, clarify) +apply (frule finite_UnionD) +apply (rotate_tac -1) +apply (induct rule: finite_induct) +apply simp_all +apply clarify apply (subst card_Un_disjoint) apply (auto simp add: disjoint_eq_subset_Compl) done @@ -2294,7 +2270,7 @@ subsubsection {* Relating injectivity and surjectivity *} -lemma finite_surj_inj: "finite(A) \ A <= f`A \ inj_on f A" +lemma finite_surj_inj: "finite A \ A \ f ` A \ inj_on f A" apply(rule eq_card_imp_inj_on, assumption) apply(frule finite_imageI) apply(drule (1) card_seteq) diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Fun.thy --- a/src/HOL/Fun.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Fun.thy Tue Jan 25 09:45:45 2011 +0100 @@ -546,12 +546,20 @@ apply (simp_all (no_asm_simp) add: inj_image_Compl_subset surj_Compl_image_subset) done +lemma inj_vimage_singleton: "inj f \ f -` {a} \ {THE x. f x = a}" + -- {* The inverse image of a singleton under an injective function + is included in a singleton. *} + apply (auto simp add: inj_on_def) + apply (blast intro: the_equality [symmetric]) + done + lemma (in ordered_ab_group_add) inj_uminus[simp, intro]: "inj_on uminus A" by (auto intro!: inj_onI) lemma (in linorder) strict_mono_imp_inj_on: "strict_mono f \ inj_on f A" by (auto intro!: inj_onI dest: strict_mono_eq) + subsection{*Function Updating*} definition diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/GCD.thy --- a/src/HOL/GCD.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/GCD.thy Tue Jan 25 09:45:45 2011 +0100 @@ -36,11 +36,8 @@ subsection {* GCD and LCM definitions *} class gcd = zero + one + dvd + - -fixes - gcd :: "'a \ 'a \ 'a" and - lcm :: "'a \ 'a \ 'a" - + fixes gcd :: "'a \ 'a \ 'a" + and lcm :: "'a \ 'a \ 'a" begin abbreviation @@ -186,7 +183,7 @@ and "x <= 0 \ y >= 0 \ P (lcm (-x) y)" and "x <= 0 \ y <= 0 \ P (lcm (-x) (-y))" shows "P (lcm x y)" -by (insert prems, auto simp add: lcm_neg1_int lcm_neg2_int, arith) + using assms by (auto simp add: lcm_neg1_int lcm_neg2_int) arith lemma lcm_ge_0_int [simp]: "lcm (x::int) y >= 0" by (simp add: lcm_int_def) @@ -632,13 +629,12 @@ apply (subgoal_tac "b' = b div gcd a b") apply (erule ssubst) apply (rule div_gcd_coprime_nat) - using prems - apply force + using z apply force apply (subst (1) b) using z apply force apply (subst (1) a) using z apply force -done + done lemma gcd_coprime_int: assumes z: "gcd (a::int) b \ 0" and a: "a = a' * gcd a b" and @@ -650,13 +646,12 @@ apply (subgoal_tac "b' = b div gcd a b") apply (erule ssubst) apply (rule div_gcd_coprime_int) - using prems - apply force + using z apply force apply (subst (1) b) using z apply force apply (subst (1) a) using z apply force -done + done lemma coprime_mult_nat: assumes da: "coprime (d::nat) a" and db: "coprime d b" shows "coprime d (a * b)" @@ -1192,13 +1187,13 @@ by auto moreover {assume db: "d=b" - from prems have ?thesis apply simp + with nz H have ?thesis apply simp apply (rule exI[where x = b], simp) apply (rule exI[where x = b]) by (rule exI[where x = "a - 1"], simp add: diff_mult_distrib2)} moreover {assume db: "d < b" - {assume "x=0" hence ?thesis using prems by simp } + {assume "x=0" hence ?thesis using nz H by simp } moreover {assume x0: "x \ 0" hence xp: "x > 0" by simp from db have "d \ b - 1" by simp diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/IMP/Com.thy --- a/src/HOL/IMP/Com.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/IMP/Com.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,8 +1,6 @@ (* Title: HOL/IMP/Com.thy - ID: $Id$ Author: Heiko Loetzbeyer & Robert Sandner & Tobias Nipkow, TUM - Isar Version: Gerwin Klein, 2001 - Copyright 1994 TUM + Author: Gerwin Klein *) header "Syntax of Commands" diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/IMP/Denotation.thy --- a/src/HOL/IMP/Denotation.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/IMP/Denotation.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/IMP/Denotation.thy - ID: $Id$ Author: Heiko Loetzbeyer & Robert Sandner, TUM - Copyright 1994 TUM *) header "Denotational Semantics of Commands" diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/IMP/Examples.thy --- a/src/HOL/IMP/Examples.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/IMP/Examples.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/IMP/Examples.thy - ID: $Id$ Author: David von Oheimb, TUM - Copyright 2000 TUM *) header "Examples" diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/IMP/Expr.thy --- a/src/HOL/IMP/Expr.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/IMP/Expr.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/IMP/Expr.thy - ID: $Id$ Author: Heiko Loetzbeyer & Robert Sandner & Tobias Nipkow, TUM - Copyright 1994 TUM *) header "Expressions" diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/IMP/Hoare.thy --- a/src/HOL/IMP/Hoare.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/IMP/Hoare.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/IMP/Hoare.thy - ID: $Id$ Author: Tobias Nipkow - Copyright 1995 TUM *) header "Inductive Definition of Hoare Logic" diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/IMP/Hoare_Den.thy --- a/src/HOL/IMP/Hoare_Den.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/IMP/Hoare_Den.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/IMP/Hoare_Def.thy - ID: $Id$ Author: Tobias Nipkow *) diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/IMP/Hoare_Op.thy --- a/src/HOL/IMP/Hoare_Op.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/IMP/Hoare_Op.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/IMP/Hoare_Op.thy - ID: $Id$ Author: Tobias Nipkow *) diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/IMP/ROOT.ML --- a/src/HOL/IMP/ROOT.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/IMP/ROOT.ML Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/IMP/ROOT.ML - ID: $Id$ Author: Heiko Loetzbeyer, Robert Sandner, Tobias Nipkow, David von Oheimb - Copyright 1995 TUM Caveat: HOLCF/IMP depends on HOL/IMP *) diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/IMP/VC.thy --- a/src/HOL/IMP/VC.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/IMP/VC.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/IMP/VC.thy - ID: $Id$ Author: Tobias Nipkow - Copyright 1996 TUM acom: annotated commands vc: verification-conditions diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/IMPP/Com.thy --- a/src/HOL/IMPP/Com.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/IMPP/Com.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/IMPP/Com.thy - ID: $Id$ Author: David von Oheimb (based on a theory by Tobias Nipkow et al), TUM - Copyright 1999 TUM *) header {* Semantics of arithmetic and boolean expressions, Syntax of commands *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/IMPP/EvenOdd.thy --- a/src/HOL/IMPP/EvenOdd.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/IMPP/EvenOdd.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/IMPP/EvenOdd.thy - ID: $Id$ - Author: David von Oheimb - Copyright 1999 TUM + Author: David von Oheimb, TUM *) header {* Example of mutually recursive procedures verified with Hoare logic *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/IMPP/Misc.thy --- a/src/HOL/IMPP/Misc.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/IMPP/Misc.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/IMPP/Misc.thy - ID: $Id$ - Author: David von Oheimb - Copyright 1999 TUM + Author: David von Oheimb, TUM *) header {* Several examples for Hoare logic *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/IMPP/Natural.thy --- a/src/HOL/IMPP/Natural.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/IMPP/Natural.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/IMPP/Natural.thy - ID: $Id$ Author: David von Oheimb (based on a theory by Tobias Nipkow et al), TUM - Copyright 1999 TUM *) header {* Natural semantics of commands *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Imperative_HOL/ex/Linked_Lists.thy --- a/src/HOL/Imperative_HOL/ex/Linked_Lists.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Imperative_HOL/ex/Linked_Lists.thy Tue Jan 25 09:45:45 2011 +0100 @@ -5,7 +5,7 @@ header {* Linked Lists by ML references *} theory Linked_Lists -imports Imperative_HOL Code_Integer +imports "../Imperative_HOL" Code_Integer begin section {* Definition of Linked Lists *} @@ -371,13 +371,12 @@ assumes "Ref.get h1 p = Node x pn" assumes "refs_of' (Ref.set p (Node x r1) h1) p rs" obtains r1s where "rs = (p#r1s)" and "refs_of' h1 r1 r1s" -using assms proof - from assms refs_of'_distinct[OF assms(2)] have "\ r1s. rs = (p # r1s) \ refs_of' h1 r1 r1s" apply - unfolding refs_of'_def'[of _ p] apply (auto, frule refs_of_set_ref2) by (auto dest: Ref.noteq_sym) - with prems show thesis by auto + with assms that show thesis by auto qed section {* Proving make_llist and traverse correct *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Import/Generate-HOL/GenHOL4Base.thy --- a/src/HOL/Import/Generate-HOL/GenHOL4Base.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Import/Generate-HOL/GenHOL4Base.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,6 +1,5 @@ (* Title: HOL/Import/Generate-HOL/GenHOL4Base.thy - ID: $Id$ - Author: Sebastian Skalberg (TU Muenchen) + Author: Sebastian Skalberg, TU Muenchen *) theory GenHOL4Base imports "../HOL4Compat" "../HOL4Syntax" begin; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Import/Generate-HOL/GenHOL4Prob.thy --- a/src/HOL/Import/Generate-HOL/GenHOL4Prob.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Import/Generate-HOL/GenHOL4Prob.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,6 +1,5 @@ (* Title: HOL/Import/Generate-HOL/GenHOL4Prob.thy - ID: $Id$ - Author: Sebastian Skalberg (TU Muenchen) + Author: Sebastian Skalberg, TU Muenchen *) theory GenHOL4Prob imports GenHOL4Real begin diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Import/Generate-HOL/GenHOL4Vec.thy --- a/src/HOL/Import/Generate-HOL/GenHOL4Vec.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Import/Generate-HOL/GenHOL4Vec.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,6 +1,5 @@ (* Title: HOL/Import/Generate-HOL/GenHOL4Vec.thy - ID: $Id$ - Author: Sebastian Skalberg (TU Muenchen) + Author: Sebastian Skalberg, TU Muenchen *) theory GenHOL4Vec imports GenHOL4Base begin diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Import/Generate-HOL/GenHOL4Word32.thy --- a/src/HOL/Import/Generate-HOL/GenHOL4Word32.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Import/Generate-HOL/GenHOL4Word32.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,6 +1,5 @@ (* Title: HOL/Import/Generate-HOL/GenHOL4Word32.thy - ID: $Id$ - Author: Sebastian Skalberg (TU Muenchen) + Author: Sebastian Skalberg, TU Muenchen *) theory GenHOL4Word32 imports GenHOL4Base begin; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Import/Generate-HOL/ROOT.ML --- a/src/HOL/Import/Generate-HOL/ROOT.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Import/Generate-HOL/ROOT.ML Tue Jan 25 09:45:45 2011 +0100 @@ -1,8 +1,3 @@ -(* Title: HOL/Import/Generate-HOL/ROOT.ML - ID: $Id$ - Author: Sebastian Skalberg (TU Muenchen) -*) - use_thy "GenHOL4Prob"; use_thy "GenHOL4Vec"; use_thy "GenHOL4Word32"; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Import/Generate-HOLLight/GenHOLLight.thy --- a/src/HOL/Import/Generate-HOLLight/GenHOLLight.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Import/Generate-HOLLight/GenHOLLight.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,6 +1,5 @@ (* Title: HOL/Import/Generate-HOLLight/GenHOLLight.thy - ID: $Id$ - Author: Steven Obua (TU Muenchen) + Author: Steven Obua, TU Muenchen *) theory GenHOLLight imports "../HOLLightCompat" "../HOL4Syntax" begin; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Import/Generate-HOLLight/ROOT.ML --- a/src/HOL/Import/Generate-HOLLight/ROOT.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Import/Generate-HOLLight/ROOT.ML Tue Jan 25 09:45:45 2011 +0100 @@ -1,6 +1,1 @@ -(* Title: HOL/Import/Generate-HOLLight/ROOT.ML - ID: $Id$ - Author: Steven Obua and Sebastian Skalberg (TU Muenchen) -*) - use_thy "GenHOLLight"; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Import/HOL/HOL4.thy --- a/src/HOL/Import/HOL/HOL4.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Import/HOL/HOL4.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,6 +1,5 @@ (* Title: HOL/Import/HOL/HOL4.thy - ID: $Id$ - Author: Sebastian Skalberg (TU Muenchen) + Author: Sebastian Skalberg, TU Muenchen *) theory HOL4 imports HOL4Vec HOL4Word32 HOL4Real begin diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Import/HOL4Compat.thy --- a/src/HOL/Import/HOL4Compat.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Import/HOL4Compat.thy Tue Jan 25 09:45:45 2011 +0100 @@ -64,10 +64,10 @@ by simp lemma sum_case_def: "(ALL f g x. sum_case f g (Inl x) = f x) & (ALL f g y. sum_case f g (Inr y) = g y)" - by simp; + by simp lemma one: "ALL v. v = ()" - by simp; + by simp lemma option_case_def: "(!u f. option_case u f None = u) & (!u f x. option_case u f (Some x) = f x)" by simp @@ -103,7 +103,7 @@ by (simp add: map_pair_def split_def) lemma pair_case_def: "split = split" - ..; + .. lemma LESS_OR_EQ: "m <= (n::nat) = (m < n | m = n)" by auto @@ -128,12 +128,12 @@ lemma LESS_DEF: "m < n = (? P. (!n. P (Suc n) --> P n) & P m & ~P n)" proof safe - assume "m < n" + assume 1: "m < n" def P == "%n. n <= m" have "(!n. P (Suc n) \ P n) & P m & ~P n" proof (auto simp add: P_def) assume "n <= m" - from prems + with 1 show False by auto qed @@ -187,7 +187,7 @@ show "m < n" .. qed -qed; +qed definition FUNPOW :: "('a => 'a) => nat => 'a => 'a" where "FUNPOW f n == f ^^ n" @@ -242,10 +242,10 @@ by auto lemma num_case_def: "(!b f. nat_case b f 0 = b) & (!b f n. nat_case b f (Suc n) = f n)" - by simp; + by simp lemma divides_def: "(a::nat) dvd b = (? q. b = q * a)" - by (auto simp add: dvd_def); + by (auto simp add: dvd_def) lemma list_case_def: "(!v f. list_case v f [] = v) & (!v f a0 a1. list_case v f (a0#a1) = f a0 a1)" by simp @@ -263,21 +263,21 @@ (list_case v f M = list_case v' f' M')" proof clarify fix M M' v f - assume "M' = [] \ v = v'" - and "!a0 a1. M' = a0 # a1 \ f a0 a1 = f' a0 a1" + assume 1: "M' = [] \ v = v'" + and 2: "!a0 a1. M' = a0 # a1 \ f a0 a1 = f' a0 a1" show "list_case v f M' = list_case v' f' M'" proof (rule List.list.case_cong) show "M' = M'" .. next assume "M' = []" - with prems + with 1 2 show "v = v'" by auto next fix a0 a1 assume "M' = a0 # a1" - with prems + with 1 2 show "f a0 a1 = f' a0 a1" by auto qed @@ -302,14 +302,14 @@ by auto next fix fn1 fn2 - assume "ALL h t. fn1 (h # t) = f (fn1 t) h t" - assume "ALL h t. fn2 (h # t) = f (fn2 t) h t" - assume "fn2 [] = fn1 []" + assume 1: "ALL h t. fn1 (h # t) = f (fn1 t) h t" + assume 2: "ALL h t. fn2 (h # t) = f (fn2 t) h t" + assume 3: "fn2 [] = fn1 []" show "fn1 = fn2" proof fix xs show "fn1 xs = fn2 xs" - by (induct xs,simp_all add: prems) + by (induct xs) (simp_all add: 1 2 3) qed qed @@ -411,7 +411,7 @@ by (simp add: Let_def) lemma REVERSE: "(rev [] = []) & (!h t. rev (h#t) = (rev t) @ [h])" - by simp; + by simp lemma REAL_SUP_ALLPOS: "\ ALL x. P (x::real) \ 0 < x ; EX x. P x; EX z. ALL x. P x \ x < z \ \ EX s. ALL y. (EX x. P x & y < x) = (y < s)" proof safe @@ -424,12 +424,11 @@ show "ALL x : Collect P. 0 < x" proof safe fix x - assume "P x" + assume P: "P x" from allx have "P x \ 0 < x" .. - thus "0 < x" - by (simp add: prems) + with P show "0 < x" by simp qed next from px @@ -461,7 +460,7 @@ by simp lemma REAL_LT_TOTAL: "((x::real) = y) | x < y | y < x" - by auto; + by auto lemma [hol4rew]: "real (0::nat) = 0" by simp diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Import/HOL4Setup.thy --- a/src/HOL/Import/HOL4Setup.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Import/HOL4Setup.thy Tue Jan 25 09:45:45 2011 +0100 @@ -90,11 +90,11 @@ have ed: "TYPE_DEFINITION P Rep" proof (auto simp add: TYPE_DEFINITION) fix x y - assume "Rep x = Rep y" + assume b: "Rep x = Rep y" from td have "x = Abs (Rep x)" by auto also have "Abs (Rep x) = Abs (Rep y)" - by (simp add: prems) + by (simp add: b) also from td have "Abs (Rep y) = y" by auto finally show "x = y" . diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Import/HOL4Syntax.thy --- a/src/HOL/Import/HOL4Syntax.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Import/HOL4Syntax.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,10 +1,11 @@ (* Title: HOL/Import/HOL4Syntax.thy - ID: $Id$ - Author: Sebastian Skalberg (TU Muenchen) + Author: Sebastian Skalberg, TU Muenchen *) -theory HOL4Syntax imports HOL4Setup - uses "import_syntax.ML" begin +theory HOL4Syntax +imports HOL4Setup +uses "import_syntax.ML" +begin ML {* HOL4ImportSyntax.setup() *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Import/HOLLight/ROOT.ML --- a/src/HOL/Import/HOLLight/ROOT.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Import/HOLLight/ROOT.ML Tue Jan 25 09:45:45 2011 +0100 @@ -1,5 +1,1 @@ -(* Title: HOL/Import/HOLLight/ROOT.ML - ID: $Id$ -*) - use_thy "HOLLight"; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Import/HOLLightCompat.thy --- a/src/HOL/Import/HOLLightCompat.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Import/HOLLightCompat.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,6 +1,5 @@ (* Title: HOL/Import/HOLLightCompat.thy - ID: $Id$ - Author: Steven Obua and Sebastian Skalberg (TU Muenchen) + Author: Steven Obua and Sebastian Skalberg, TU Muenchen *) theory HOLLightCompat imports HOL4Setup HOL4Compat Divides Primes Real begin diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Import/MakeEqual.thy --- a/src/HOL/Import/MakeEqual.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Import/MakeEqual.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,6 +1,5 @@ (* Title: HOL/Import/MakeEqual.thy - ID: $Id$ - Author: Sebastian Skalberg (TU Muenchen) + Author: Sebastian Skalberg, TU Muenchen *) theory MakeEqual imports Main diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Import/mono_scan.ML --- a/src/HOL/Import/mono_scan.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Import/mono_scan.ML Tue Jan 25 09:45:45 2011 +0100 @@ -1,8 +1,7 @@ (* Title: HOL/Import/mono_scan.ML - ID: $Id$ Author: Steven Obua, TU Muenchen - Monomorphic scanner combinators for monomorphic sequences. +Monomorphic scanner combinators for monomorphic sequences. *) signature MONO_SCANNER = diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Import/mono_seq.ML --- a/src/HOL/Import/mono_seq.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Import/mono_seq.ML Tue Jan 25 09:45:45 2011 +0100 @@ -1,8 +1,7 @@ (* Title: HOL/Import/mono_seq.ML - ID: $Id$ Author: Steven Obua, TU Muenchen - Monomorphic sequences. +Monomorphic sequences. *) (* The trouble is that signature / structures cannot depend on type variable parameters ... *) diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/IsaMakefile Tue Jan 25 09:45:45 2011 +0100 @@ -17,6 +17,7 @@ HOL-Nominal \ HOL-Probability \ HOL-Proofs \ + HOL-SPARK \ HOL-Word \ HOL4 \ HOLCF \ @@ -71,6 +72,7 @@ HOL-Proofs-Extraction \ HOL-Proofs-Lambda \ HOL-SET_Protocol \ + HOL-SPARK-Examples \ HOL-Word-SMT_Examples \ HOL-Statespace \ HOL-Subst \ @@ -1031,19 +1033,20 @@ ex/HarmonicSeries.thy ex/Hebrew.thy ex/Hex_Bin_Examples.thy \ ex/Higher_Order_Logic.thy ex/Iff_Oracle.thy ex/Induction_Schema.thy \ ex/InductiveInvariant.thy ex/InductiveInvariant_examples.thy \ - ex/Intuitionistic.thy ex/Lagrange.thy \ - ex/List_to_Set_Comprehension_Examples.thy ex/LocaleTest2.thy ex/MT.thy \ - ex/MergeSort.thy ex/Meson_Test.thy ex/MonoidGroup.thy \ + ex/Interpretation_with_Defs.thy ex/Intuitionistic.thy ex/Lagrange.thy \ + ex/List_to_Set_Comprehension_Examples.thy ex/LocaleTest2.thy \ + ex/MT.thy ex/MergeSort.thy ex/Meson_Test.thy ex/MonoidGroup.thy \ ex/Multiquote.thy ex/NatSum.thy ex/Normalization_by_Evaluation.thy \ ex/Numeral.thy ex/PER.thy ex/PresburgerEx.thy ex/Primrec.thy \ ex/Quickcheck_Examples.thy ex/Quickcheck_Lattice_Examples.thy \ ex/Quicksort.thy ex/ROOT.ML ex/Recdefs.thy ex/Records.thy \ ex/ReflectionEx.thy ex/Refute_Examples.thy ex/SAT_Examples.thy \ - ex/SVC_Oracle.thy ex/Serbian.thy ex/Sqrt.thy ex/Sqrt_Script.thy \ - ex/Sudoku.thy ex/Tarski.thy ex/Termination.thy ex/Transfer_Ex.thy \ - ex/Tree23.thy ex/Unification.thy ex/While_Combinator_Example.thy \ - ex/document/root.bib ex/document/root.tex ex/set.thy ex/svc_funcs.ML \ - ex/svc_test.thy + ex/SVC_Oracle.thy ex/Serbian.thy ex/Set_Algebras.thy ex/Sqrt.thy \ + ex/Sqrt_Script.thy ex/Sudoku.thy ex/Tarski.thy ex/Termination.thy \ + ex/Transfer_Ex.thy ex/Tree23.thy ex/Unification.thy \ + ex/While_Combinator_Example.thy ex/document/root.bib \ + ex/document/root.tex ex/set.thy ex/svc_funcs.ML ex/svc_test.thy \ + ../Tools/interpretation_with_defs.ML @$(ISABELLE_TOOL) usedir $(OUT)/HOL ex @@ -1199,7 +1202,6 @@ Nominal/nominal_permeq.ML \ Nominal/nominal_primrec.ML \ Nominal/nominal_thmdecls.ML \ - Nominal/old_primrec.ML \ Library/Infinite_Set.thy @cd Nominal; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL HOL-Nominal @@ -1344,6 +1346,69 @@ @cd Boogie; $(ISABELLE_TOOL) usedir $(OUT)/HOL-Boogie Examples +## HOL-SPARK + +HOL-SPARK: HOL-Word $(OUT)/HOL-SPARK + +$(OUT)/HOL-SPARK: $(OUT)/HOL-Word SPARK/ROOT.ML \ + SPARK/SPARK.thy SPARK/SPARK_Setup.thy \ + SPARK/Tools/fdl_lexer.ML SPARK/Tools/fdl_parser.ML \ + SPARK/Tools/spark_commands.ML SPARK/Tools/spark_vcs.ML + @cd SPARK; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL-Word HOL-SPARK + + +## HOL-SPARK-Examples + +HOL-SPARK-Examples: HOL-SPARK $(LOG)/HOL-SPARK-Examples.gz + +$(LOG)/HOL-SPARK-Examples.gz: $(OUT)/HOL-SPARK \ + SPARK/Examples/ROOT.ML \ + SPARK/Examples/Gcd/Greatest_Common_Divisor.thy \ + SPARK/Examples/Gcd/greatest_common_divisor/g_c_d.fdl \ + SPARK/Examples/Gcd/greatest_common_divisor/g_c_d.rls \ + SPARK/Examples/Gcd/greatest_common_divisor/g_c_d.siv \ + SPARK/Examples/Liseq/Longest_Increasing_Subsequence.thy \ + SPARK/Examples/Liseq/liseq/liseq_length.fdl \ + SPARK/Examples/Liseq/liseq/liseq_length.rls \ + SPARK/Examples/Liseq/liseq/liseq_length.siv \ + SPARK/Examples/RIPEMD-160/F.thy SPARK/Examples/RIPEMD-160/Hash.thy \ + SPARK/Examples/RIPEMD-160/K_L.thy SPARK/Examples/RIPEMD-160/K_R.thy \ + SPARK/Examples/RIPEMD-160/R_L.thy \ + SPARK/Examples/RIPEMD-160/RMD_Lemmas.thy \ + SPARK/Examples/RIPEMD-160/RMD_Specification.thy \ + SPARK/Examples/RIPEMD-160/RMD.thy SPARK/Examples/RIPEMD-160/Round.thy \ + SPARK/Examples/RIPEMD-160/R_R.thy SPARK/Examples/RIPEMD-160/S_L.thy \ + SPARK/Examples/RIPEMD-160/S_R.thy \ + SPARK/Examples/RIPEMD-160/rmd/f.fdl \ + SPARK/Examples/RIPEMD-160/rmd/f.rls \ + SPARK/Examples/RIPEMD-160/rmd/f.siv \ + SPARK/Examples/RIPEMD-160/rmd/hash.fdl \ + SPARK/Examples/RIPEMD-160/rmd/hash.rls \ + SPARK/Examples/RIPEMD-160/rmd/hash.siv \ + SPARK/Examples/RIPEMD-160/rmd/k_l.fdl \ + SPARK/Examples/RIPEMD-160/rmd/k_l.rls \ + SPARK/Examples/RIPEMD-160/rmd/k_l.siv \ + SPARK/Examples/RIPEMD-160/rmd/k_r.fdl \ + SPARK/Examples/RIPEMD-160/rmd/k_r.rls \ + SPARK/Examples/RIPEMD-160/rmd/k_r.siv \ + SPARK/Examples/RIPEMD-160/rmd/r_l.fdl \ + SPARK/Examples/RIPEMD-160/rmd/r_l.rls \ + SPARK/Examples/RIPEMD-160/rmd/r_l.siv \ + SPARK/Examples/RIPEMD-160/rmd/round.fdl \ + SPARK/Examples/RIPEMD-160/rmd/round.rls \ + SPARK/Examples/RIPEMD-160/rmd/round.siv \ + SPARK/Examples/RIPEMD-160/rmd/r_r.fdl \ + SPARK/Examples/RIPEMD-160/rmd/r_r.rls \ + SPARK/Examples/RIPEMD-160/rmd/r_r.siv \ + SPARK/Examples/RIPEMD-160/rmd/s_l.fdl \ + SPARK/Examples/RIPEMD-160/rmd/s_l.rls \ + SPARK/Examples/RIPEMD-160/rmd/s_l.siv \ + SPARK/Examples/RIPEMD-160/rmd/s_r.fdl \ + SPARK/Examples/RIPEMD-160/rmd/s_r.rls \ + SPARK/Examples/RIPEMD-160/rmd/s_r.siv + @cd SPARK; $(ISABELLE_TOOL) usedir $(OUT)/HOL-SPARK Examples + + ## HOL-Mutabelle HOL-Mutabelle: HOL $(LOG)/HOL-Mutabelle.gz @@ -1649,6 +1714,7 @@ $(LOG)/HOL-Proofs-Extraction.gz \ $(LOG)/HOL-Proofs-Lambda.gz $(LOG)/HOL-SET_Protocol.gz \ $(LOG)/HOL-Word-SMT_Examples.gz \ + $(LOG)/HOL-SPARK.gz $(LOG)/HOL-SPARK-Examples.gz \ $(LOG)/HOL-Statespace.gz $(LOG)/HOL-Subst.gz \ $(LOG)/HOL-UNITY.gz $(LOG)/HOL-Unix.gz \ $(LOG)/HOL-Word-Examples.gz $(LOG)/HOL-Word.gz \ @@ -1659,6 +1725,7 @@ $(OUT)/HOL-Main $(OUT)/HOL-Multivariate_Analysis \ $(OUT)/HOL-NSA $(OUT)/HOL-Nominal $(OUT)/HOL-Plain \ $(OUT)/HOL-Probability $(OUT)/HOL-Proofs \ + $(OUT)/HOL-SPARK \ $(OUT)/HOL-Word $(OUT)/HOL4 $(OUT)/TLA \ $(OUT)/HOLCF $(LOG)/HOLCF.gz $(LOG)/HOLCF-IMP.gz \ $(LOG)/HOLCF-ex.gz $(LOG)/HOLCF-FOCUS.gz $(OUT)/IOA \ diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Lim.thy --- a/src/HOL/Lim.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Lim.thy Tue Jan 25 09:45:45 2011 +0100 @@ -653,7 +653,7 @@ moreover have "\n. ?F n \ a" by (rule allI) (rule F1) - moreover from prems have "\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L" by simp + moreover note `\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L` ultimately have "(\n. X (?F n)) ----> L" by simp moreover have "\ ((\n. X (?F n)) ----> L)" diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Ln.thy --- a/src/HOL/Ln.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Ln.thy Tue Jan 25 09:45:45 2011 +0100 @@ -71,7 +71,7 @@ qed moreover have "x ^ (Suc n + 2) <= x ^ (n + 2)" apply (simp add: mult_compare_simps) - apply (simp add: prems) + apply (simp add: assms) apply (subgoal_tac "0 <= x * (x * x^n)") apply force apply (rule mult_nonneg_nonneg, rule a)+ @@ -91,7 +91,7 @@ by simp also have "... <= 1 / 2 * (x ^ 2 / 2 * (1 / 2) ^ n)" apply (rule mult_left_mono) - apply (rule prems) + apply (rule c) apply simp done also have "... = x ^ 2 / 2 * (1 / 2 * (1 / 2) ^ n)" @@ -129,7 +129,7 @@ have "suminf (%n. inverse(fact (n+2)) * (x ^ (n+2))) <= suminf (%n. (x^2/2) * ((1/2)^n))" apply (rule summable_le) - apply (auto simp only: aux1 prems) + apply (auto simp only: aux1 a b) apply (rule exp_tail_after_first_two_terms_summable) by (rule sums_summable, rule aux2) also have "... = x^2" @@ -155,14 +155,14 @@ apply (rule divide_left_mono) apply (auto simp add: exp_ge_add_one_self_aux) apply (rule add_nonneg_nonneg) - apply (insert prems, auto) + using a apply auto apply (rule mult_pos_pos) apply auto apply (rule add_pos_nonneg) apply auto done also from a have "... <= 1 + x" - by(simp add:field_simps zero_compare_simps) + by (simp add: field_simps zero_compare_simps) finally show ?thesis . qed @@ -192,14 +192,14 @@ finally have "(1 - x) * (1 + x + x ^ 2) <= 1" . moreover have "0 < 1 + x + x^2" apply (rule add_pos_nonneg) - apply (insert a, auto) + using a apply auto done ultimately have "1 - x <= 1 / (1 + x + x^2)" by (elim mult_imp_le_div_pos) also have "... <= 1 / exp x" apply (rule divide_left_mono) apply (rule exp_bound, rule a) - apply (insert prems, auto) + using a b apply auto apply (rule mult_pos_pos) apply (rule add_pos_nonneg) apply auto @@ -256,10 +256,10 @@ also have "- (x / (1 - x)) = -x / (1 - x)" by auto finally have d: "- x / (1 - x) <= ln (1 - x)" . - have "0 < 1 - x" using prems by simp + have "0 < 1 - x" using a b by simp hence e: "-x - 2 * x^2 <= - x / (1 - x)" - using mult_right_le_one_le[of "x*x" "2*x"] prems - by(simp add:field_simps power2_eq_square) + using mult_right_le_one_le[of "x*x" "2*x"] a b + by (simp add:field_simps power2_eq_square) from e d show "- x - 2 * x^2 <= ln (1 - x)" by (rule order_trans) qed @@ -292,7 +292,7 @@ "0 <= x ==> x <= 1 ==> abs(ln (1 + x) - x) <= x^2" proof - assume x: "0 <= x" - assume "x <= 1" + assume x1: "x <= 1" from x have "ln (1 + x) <= x" by (rule ln_add_one_self_le_self) then have "ln (1 + x) - x <= 0" @@ -303,7 +303,7 @@ by simp also have "... <= x^2" proof - - from prems have "x - x^2 <= ln (1 + x)" + from x x1 have "x - x^2 <= ln (1 + x)" by (intro ln_one_plus_pos_lower_bound) thus ?thesis by simp @@ -314,19 +314,19 @@ lemma abs_ln_one_plus_x_minus_x_bound_nonpos: "-(1 / 2) <= x ==> x <= 0 ==> abs(ln (1 + x) - x) <= 2 * x^2" proof - - assume "-(1 / 2) <= x" - assume "x <= 0" + assume a: "-(1 / 2) <= x" + assume b: "x <= 0" have "abs(ln (1 + x) - x) = x - ln(1 - (-x))" apply (subst abs_of_nonpos) apply simp apply (rule ln_add_one_self_le_self2) - apply (insert prems, auto) + using a apply auto done also have "... <= 2 * x^2" apply (subgoal_tac "- (-x) - 2 * (-x)^2 <= ln (1 - (-x))") apply (simp add: algebra_simps) apply (rule ln_one_minus_pos_lower_bound) - apply (insert prems, auto) + using a b apply auto done finally show ?thesis . qed @@ -343,9 +343,9 @@ lemma ln_x_over_x_mono: "exp 1 <= x ==> x <= y ==> (ln y / y) <= (ln x / x)" proof - - assume "exp 1 <= x" and "x <= y" + assume x: "exp 1 <= x" "x <= y" have a: "0 < x" and b: "0 < y" - apply (insert prems) + apply (insert x) apply (subgoal_tac "0 < exp (1::real)") apply arith apply auto @@ -361,12 +361,12 @@ done also have "y / x = (x + (y - x)) / x" by simp - also have "... = 1 + (y - x) / x" using a prems by(simp add:field_simps) + also have "... = 1 + (y - x) / x" using x a by (simp add: field_simps) also have "x * ln(1 + (y - x) / x) <= x * ((y - x) / x)" apply (rule mult_left_mono) apply (rule ln_add_one_self_le_self) apply (rule divide_nonneg_pos) - apply (insert prems a, simp_all) + using x a apply simp_all done also have "... = y - x" using a by simp also have "... = (y - x) * ln (exp 1)" by simp @@ -375,16 +375,16 @@ apply (subst ln_le_cancel_iff) apply force apply (rule a) - apply (rule prems) - apply (insert prems, simp) + apply (rule x) + using x apply simp done also have "... = y * ln x - x * ln x" by (rule left_diff_distrib) finally have "x * ln y <= y * ln x" by arith - then have "ln y <= (y * ln x) / x" using a by(simp add:field_simps) - also have "... = y * (ln x / x)" by simp - finally show ?thesis using b by(simp add:field_simps) + then have "ln y <= (y * ln x) / x" using a by (simp add: field_simps) + also have "... = y * (ln x / x)" by simp + finally show ?thesis using b by (simp add: field_simps) qed end diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Log.thy --- a/src/HOL/Log.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Log.thy Tue Jan 25 09:45:45 2011 +0100 @@ -251,10 +251,11 @@ apply (erule order_less_imp_le) done -lemma ln_powr_bound2: "1 < x ==> 0 < a ==> (ln x) powr a <= (a powr a) * x" +lemma ln_powr_bound2: + assumes "1 < x" and "0 < a" + shows "(ln x) powr a <= (a powr a) * x" proof - - assume "1 < x" and "0 < a" - then have "ln x <= (x powr (1 / a)) / (1 / a)" + from assms have "ln x <= (x powr (1 / a)) / (1 / a)" apply (intro ln_powr_bound) apply (erule order_less_imp_le) apply (rule divide_pos_pos) @@ -264,14 +265,14 @@ by simp finally have "(ln x) powr a <= (a * (x powr (1 / a))) powr a" apply (intro powr_mono2) - apply (rule order_less_imp_le, rule prems) + apply (rule order_less_imp_le, rule assms) apply (rule ln_gt_zero) - apply (rule prems) + apply (rule assms) apply assumption done also have "... = (a powr a) * ((x powr (1 / a)) powr a)" apply (rule powr_mult) - apply (rule prems) + apply (rule assms) apply (rule powr_gt_zero) done also have "(x powr (1 / a)) powr a = x powr ((1 / a) * a)" @@ -279,35 +280,37 @@ also have "... = x" apply simp apply (subgoal_tac "a ~= 0") - apply (insert prems, auto) + using assms apply auto done finally show ?thesis . qed -lemma LIMSEQ_neg_powr: "0 < s ==> (%x. (real x) powr - s) ----> 0" +lemma LIMSEQ_neg_powr: + assumes s: "0 < s" + shows "(%x. (real x) powr - s) ----> 0" apply (unfold LIMSEQ_iff) apply clarsimp apply (rule_tac x = "natfloor(r powr (1 / - s)) + 1" in exI) apply clarify - proof - - fix r fix n - assume "0 < s" and "0 < r" and "natfloor (r powr (1 / - s)) + 1 <= n" - have "r powr (1 / - s) < real(natfloor(r powr (1 / - s))) + 1" - by (rule real_natfloor_add_one_gt) - also have "... = real(natfloor(r powr (1 / -s)) + 1)" - by simp - also have "... <= real n" - apply (subst real_of_nat_le_iff) - apply (rule prems) - done - finally have "r powr (1 / - s) < real n". - then have "real n powr (- s) < (r powr (1 / - s)) powr - s" - apply (intro powr_less_mono2_neg) - apply (auto simp add: prems) - done - also have "... = r" - by (simp add: powr_powr prems less_imp_neq [THEN not_sym]) - finally show "real n powr - s < r" . - qed +proof - + fix r fix n + assume r: "0 < r" and n: "natfloor (r powr (1 / - s)) + 1 <= n" + have "r powr (1 / - s) < real(natfloor(r powr (1 / - s))) + 1" + by (rule real_natfloor_add_one_gt) + also have "... = real(natfloor(r powr (1 / -s)) + 1)" + by simp + also have "... <= real n" + apply (subst real_of_nat_le_iff) + apply (rule n) + done + finally have "r powr (1 / - s) < real n". + then have "real n powr (- s) < (r powr (1 / - s)) powr - s" + apply (intro powr_less_mono2_neg) + apply (auto simp add: s) + done + also have "... = r" + by (simp add: powr_powr s r less_imp_neq [THEN not_sym]) + finally show "real n powr - s < r" . +qed end diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Map.thy --- a/src/HOL/Map.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Map.thy Tue Jan 25 09:45:45 2011 +0100 @@ -111,7 +111,7 @@ assumes "m(a\x) = n(a\y)" shows "x = y" proof - - from prems have "(m(a\x)) a = (n(a\y)) a" by simp + from assms have "(m(a\x)) a = (n(a\y)) a" by simp then show ?thesis by simp qed diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Matrix/ComputeFloat.thy --- a/src/HOL/Matrix/ComputeFloat.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Matrix/ComputeFloat.thy Tue Jan 25 09:45:45 2011 +0100 @@ -57,7 +57,7 @@ show ?case by simp next case (Suc m) - show ?case by (auto simp add: algebra_simps pow2_add1 prems) + show ?case by (auto simp add: algebra_simps pow2_add1 1 Suc) qed next case (2 n) @@ -88,7 +88,7 @@ apply (subst pow2_neg[of "int m - a + 1"]) apply (subst pow2_neg[of "int m + 1"]) apply auto - apply (insert prems) + apply (insert Suc) apply (auto simp add: algebra_simps) done qed @@ -147,8 +147,8 @@ assumes "real_is_int a" "real_is_int b" shows "(int_of_real (a*b)) = (int_of_real a) * (int_of_real b)" proof - - from prems have a: "?! (a'::int). real a' = a" by (rule_tac real_is_int_rep, auto) - from prems have b: "?! (b'::int). real b' = b" by (rule_tac real_is_int_rep, auto) + from assms have a: "?! (a'::int). real a' = a" by (rule_tac real_is_int_rep, auto) + from assms have b: "?! (b'::int). real b' = b" by (rule_tac real_is_int_rep, auto) from a obtain a'::int where a':"a = real a'" by auto from b obtain b'::int where b':"b = real b'" by auto have r: "real a' * real b' = real (a' * b')" by auto @@ -286,16 +286,16 @@ show ?case proof cases assume u: "u \ 0 \ even u" - with prems have ind: "float (u div 2, v + 1) = float (norm_float (u div 2) (v + 1))" by auto + with 1 have ind: "float (u div 2, v + 1) = float (norm_float (u div 2) (v + 1))" by auto with u have "float (u,v) = float (u div 2, v+1)" by (simp add: float_transfer_even) then show ?thesis apply (subst norm_float.simps) apply (simp add: ind) done next - assume "~(u \ 0 \ even u)" - then show ?thesis - by (simp add: prems float_def) + assume nu: "~(u \ 0 \ even u)" + show ?thesis + by (simp add: nu float_def) qed qed } diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Matrix/LP.thy --- a/src/HOL/Matrix/LP.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Matrix/LP.thy Tue Jan 25 09:45:45 2011 +0100 @@ -12,7 +12,7 @@ "c <= d" shows "a <= b + d" apply (rule_tac order_trans[where y = "b+c"]) - apply (simp_all add: prems) + apply (simp_all add: assms) done lemma linprog_dual_estimate: @@ -26,8 +26,8 @@ shows "c * x \ y * b' + (y * \A + abs (y * A' - c') + \c) * r" proof - - from prems have 1: "y * b <= y * b'" by (simp add: mult_left_mono) - from prems have 2: "y * (A * x) <= y * b" by (simp add: mult_left_mono) + from assms have 1: "y * b <= y * b'" by (simp add: mult_left_mono) + from assms have 2: "y * (A * x) <= y * b" by (simp add: mult_left_mono) have 3: "y * (A * x) = c * x + (y * (A - A') + (y * A' - c') + (c'-c)) * x" by (simp add: algebra_simps) from 1 2 3 have 4: "c * x + (y * (A - A') + (y * A' - c') + (c'-c)) * x <= y * b'" by simp have 5: "c * x <= y * b' + abs((y * (A - A') + (y * A' - c') + (c'-c)) * x)" @@ -44,23 +44,23 @@ have 11: "abs (c'-c) = abs (c-c')" by (subst 10, subst abs_minus_cancel, simp) have 12: "(abs y * abs (A-A') + abs (y*A'-c') + abs (c'-c)) * abs x <= (abs y * abs (A-A') + abs (y*A'-c') + \c) * abs x" - by (simp add: 11 prems mult_right_mono) + by (simp add: 11 assms mult_right_mono) have 13: "(abs y * abs (A-A') + abs (y*A'-c') + \c) * abs x <= (abs y * \A + abs (y*A'-c') + \c) * abs x" - by (simp add: prems mult_right_mono mult_left_mono) + by (simp add: assms mult_right_mono mult_left_mono) have r: "(abs y * \A + abs (y*A'-c') + \c) * abs x <= (abs y * \A + abs (y*A'-c') + \c) * r" apply (rule mult_left_mono) - apply (simp add: prems) + apply (simp add: assms) apply (rule_tac add_mono[of "0::'a" _ "0", simplified])+ apply (rule mult_left_mono[of "0" "\A", simplified]) apply (simp_all) - apply (rule order_trans[where y="abs (A-A')"], simp_all add: prems) - apply (rule order_trans[where y="abs (c-c')"], simp_all add: prems) + apply (rule order_trans[where y="abs (A-A')"], simp_all add: assms) + apply (rule order_trans[where y="abs (c-c')"], simp_all add: assms) done from 6 7 8 9 12 13 r have 14:" abs((y * (A - A') + (y * A' - c') + (c'-c)) * x) <=(abs y * \A + abs (y*A'-c') + \c) * r" by (simp) show ?thesis apply (rule le_add_right_mono[of _ _ "abs((y * (A - A') + (y * A' - c') + (c'-c)) * x)"]) - apply (simp_all only: 5 14[simplified abs_of_nonneg[of y, simplified prems]]) + apply (simp_all only: 5 14[simplified abs_of_nonneg[of y, simplified assms]]) done qed @@ -73,10 +73,10 @@ have "0 <= A - A1" proof - have 1: "A - A1 = A + (- A1)" by simp - show ?thesis by (simp only: 1 add_right_mono[of A1 A "-A1", simplified, simplified prems]) + show ?thesis by (simp only: 1 add_right_mono[of A1 A "-A1", simplified, simplified assms]) qed then have "abs (A-A1) = A-A1" by (rule abs_of_nonneg) - with prems show "abs (A-A1) <= (A2-A1)" by simp + with assms show "abs (A-A1) <= (A2-A1)" by simp qed lemma mult_le_prts: @@ -95,31 +95,31 @@ then have "a * b = pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b" by (simp add: algebra_simps) moreover have "pprt a * pprt b <= pprt a2 * pprt b2" - by (simp_all add: prems mult_mono) + by (simp_all add: assms mult_mono) moreover have "pprt a * nprt b <= pprt a1 * nprt b2" proof - have "pprt a * nprt b <= pprt a * nprt b2" - by (simp add: mult_left_mono prems) + by (simp add: mult_left_mono assms) moreover have "pprt a * nprt b2 <= pprt a1 * nprt b2" - by (simp add: mult_right_mono_neg prems) + by (simp add: mult_right_mono_neg assms) ultimately show ?thesis by simp qed moreover have "nprt a * pprt b <= nprt a2 * pprt b1" proof - have "nprt a * pprt b <= nprt a2 * pprt b" - by (simp add: mult_right_mono prems) + by (simp add: mult_right_mono assms) moreover have "nprt a2 * pprt b <= nprt a2 * pprt b1" - by (simp add: mult_left_mono_neg prems) + by (simp add: mult_left_mono_neg assms) ultimately show ?thesis by simp qed moreover have "nprt a * nprt b <= nprt a1 * nprt b1" proof - have "nprt a * nprt b <= nprt a * nprt b1" - by (simp add: mult_left_mono_neg prems) + by (simp add: mult_left_mono_neg assms) moreover have "nprt a * nprt b1 <= nprt a1 * nprt b1" - by (simp add: mult_right_mono_neg prems) + by (simp add: mult_right_mono_neg assms) ultimately show ?thesis by simp qed @@ -141,19 +141,19 @@ "c * x \ y * b + (let s1 = c1 - y * A2; s2 = c2 - y * A1 in pprt s2 * pprt r2 + pprt s1 * nprt r2 + nprt s2 * pprt r1 + nprt s1 * nprt r1)" (is "_ <= _ + ?C") proof - - from prems have "y * (A * x) <= y * b" by (simp add: mult_left_mono) + from assms have "y * (A * x) <= y * b" by (simp add: mult_left_mono) moreover have "y * (A * x) = c * x + (y * A - c) * x" by (simp add: algebra_simps) ultimately have "c * x + (y * A - c) * x <= y * b" by simp then have "c * x <= y * b - (y * A - c) * x" by (simp add: le_diff_eq) then have cx: "c * x <= y * b + (c - y * A) * x" by (simp add: algebra_simps) have s2: "c - y * A <= c2 - y * A1" - by (simp add: diff_minus prems add_mono mult_left_mono) + by (simp add: diff_minus assms add_mono mult_left_mono) have s1: "c1 - y * A2 <= c - y * A" - by (simp add: diff_minus prems add_mono mult_left_mono) + by (simp add: diff_minus assms add_mono mult_left_mono) have prts: "(c - y * A) * x <= ?C" apply (simp add: Let_def) apply (rule mult_le_prts) - apply (simp_all add: prems s1 s2) + apply (simp_all add: assms s1 s2) done then have "y * b + (c - y * A) * x <= y * b + ?C" by simp diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/MicroJava/Comp/NatCanonify.thy --- a/src/HOL/MicroJava/Comp/NatCanonify.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/MicroJava/Comp/NatCanonify.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/MicroJava/Comp/NatCanonify.thy - ID: $Id$ Author: Martin Strecker *) diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/MicroJava/J/JBasis.thy --- a/src/HOL/MicroJava/J/JBasis.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/MicroJava/J/JBasis.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/MicroJava/J/JBasis.thy - ID: $Id$ - Author: David von Oheimb - Copyright 1999 TU Muenchen + Author: David von Oheimb, TU Muenchen *) header {* diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/MicroJava/J/JTypeSafe.thy --- a/src/HOL/MicroJava/J/JTypeSafe.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/MicroJava/J/JTypeSafe.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/MicroJava/J/JTypeSafe.thy - ID: $Id$ - Author: David von Oheimb - Copyright 1999 Technische Universitaet Muenchen + Author: David von Oheimb, Technische Universitaet Muenchen *) header {* \isaheader{Type Safety Proof} *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/MicroJava/J/Term.thy --- a/src/HOL/MicroJava/J/Term.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/MicroJava/J/Term.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/MicroJava/J/Term.thy - ID: $Id$ - Author: David von Oheimb - Copyright 1999 Technische Universitaet Muenchen + Author: David von Oheimb, Technische Universitaet Muenchen *) header {* \isaheader{Expressions and Statements} *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/MicroJava/J/TypeRel.thy --- a/src/HOL/MicroJava/J/TypeRel.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/MicroJava/J/TypeRel.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/MicroJava/J/TypeRel.thy - ID: $Id$ - Author: David von Oheimb - Copyright 1999 Technische Universitaet Muenchen + Author: David von Oheimb, Technische Universitaet Muenchen *) header {* \isaheader{Relations between Java Types} *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/MicroJava/J/Value.thy --- a/src/HOL/MicroJava/J/Value.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/MicroJava/J/Value.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/MicroJava/J/Value.thy - ID: $Id$ - Author: David von Oheimb - Copyright 1999 Technische Universitaet Muenchen + Author: David von Oheimb, Technische Universitaet Muenchen *) header {* \isaheader{Java Values} *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/MicroJava/JVM/JVMInstructions.thy --- a/src/HOL/MicroJava/JVM/JVMInstructions.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/MicroJava/JVM/JVMInstructions.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/MicroJava/JVM/JVMInstructions.thy - ID: $Id$ - Author: Gerwin Klein - Copyright 2000 Technische Universitaet Muenchen + Author: Gerwin Klein, Technische Universitaet Muenchen *) header {* \isaheader{Instructions of the JVM} *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/MicroJava/JVM/JVMState.thy --- a/src/HOL/MicroJava/JVM/JVMState.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/MicroJava/JVM/JVMState.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/MicroJava/JVM/JVMState.thy - ID: $Id$ - Author: Cornelia Pusch, Gerwin Klein - Copyright 1999 Technische Universitaet Muenchen + Author: Cornelia Pusch, Gerwin Klein, Technische Universitaet Muenchen *) header {* diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Tue Jan 25 09:45:45 2011 +0100 @@ -12,7 +12,7 @@ declare [[smt_certificates="Integration.certs"]] declare [[smt_fixed=true]] -declare [[smt_solver=z3, smt_oracle=false]] +declare [[smt_oracle=false]] setup {* Arith_Data.add_tactic "Ferrante-Rackoff" (K FerranteRackoff.dlo_tac) *} @@ -5527,6 +5527,5 @@ declare [[smt_certificates=""]] declare [[smt_fixed=false]] -declare [[smt_solver=cvc3]] end diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/NSA/Filter.thy --- a/src/HOL/NSA/Filter.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/NSA/Filter.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,9 +1,7 @@ -(* Title : Filter.thy - ID : $Id$ - Author : Jacques D. Fleuriot - Copyright : 1998 University of Cambridge - Conversion to Isar and new proofs by Lawrence C Paulson, 2004 - Conversion to locales by Brian Huffman, 2005 +(* Title: Filter.thy + Author: Jacques D. Fleuriot, University of Cambridge + Author: Lawrence C Paulson + Author: Brian Huffman *) header {* Filters and Ultrafilters *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/NSA/HLim.thy --- a/src/HOL/NSA/HLim.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/NSA/HLim.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,8 +1,6 @@ -(* Title : HLim.thy - ID : $Id$ - Author : Jacques D. Fleuriot - Copyright : 1998 University of Cambridge - Conversion to Isar and new proofs by Lawrence C Paulson, 2004 +(* Title: HLim.thy + Author: Jacques D. Fleuriot, University of Cambridge + Author: Lawrence C Paulson *) header{* Limits and Continuity (Nonstandard) *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/NanoJava/AxSem.thy --- a/src/HOL/NanoJava/AxSem.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/NanoJava/AxSem.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/NanoJava/AxSem.thy - ID: $Id$ - Author: David von Oheimb - Copyright 2001 Technische Universitaet Muenchen + Author: David von Oheimb, Technische Universitaet Muenchen *) header "Axiomatic Semantics" diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/NanoJava/Term.thy --- a/src/HOL/NanoJava/Term.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/NanoJava/Term.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/NanoJava/Term.thy - ID: $Id$ - Author: David von Oheimb - Copyright 2001 Technische Universitaet Muenchen + Author: David von Oheimb, Technische Universitaet Muenchen *) header "Statements and expression emulations" diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/NanoJava/TypeRel.thy --- a/src/HOL/NanoJava/TypeRel.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/NanoJava/TypeRel.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/NanoJava/TypeRel.thy - ID: $Id$ - Author: David von Oheimb - Copyright 2001 Technische Universitaet Muenchen + Author: David von Oheimb, Technische Universitaet Muenchen *) header "Type relations" diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Nominal/Examples/CR.thy --- a/src/HOL/Nominal/Examples/CR.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Nominal/Examples/CR.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,5 +1,3 @@ -(* $Id$ *) - theory CR imports Lam_Funs begin diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Nominal/Examples/Lambda_mu.thy --- a/src/HOL/Nominal/Examples/Lambda_mu.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Nominal/Examples/Lambda_mu.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,5 +1,3 @@ -(* $Id$ *) - theory Lambda_mu imports "../Nominal" begin diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Nominal/Examples/Support.thy --- a/src/HOL/Nominal/Examples/Support.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Nominal/Examples/Support.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,5 +1,3 @@ -(* $Id$ *) - theory Support imports "../Nominal" begin diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Nominal/Examples/VC_Condition.thy --- a/src/HOL/Nominal/Examples/VC_Condition.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Nominal/Examples/VC_Condition.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,5 +1,3 @@ -(* $Id$ *) - theory VC_Condition imports "../Nominal" begin diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Nominal/Nominal.thy --- a/src/HOL/Nominal/Nominal.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Nominal/Nominal.thy Tue Jan 25 09:45:45 2011 +0100 @@ -10,7 +10,6 @@ ("nominal_primrec.ML") ("nominal_inductive.ML") ("nominal_inductive2.ML") - ("old_primrec.ML") begin section {* Permutations *} @@ -785,7 +784,7 @@ hence "((UNIV::'x set) - A) \ ({}::'x set)" by (force simp only:) then obtain c::"'x" where "c\((UNIV::'x set) - A)" by force then have "c\A" by simp - then show ?thesis using prems by simp + then show ?thesis .. qed text {* there always exists a fresh name for an object with finite support *} @@ -3605,7 +3604,6 @@ (***************************************) (* setup for the individial atom-kinds *) (* and nominal datatypes *) -use "old_primrec.ML" use "nominal_atoms.ML" (************************************************************) diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Nominal/nominal_atoms.ML --- a/src/HOL/Nominal/nominal_atoms.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Nominal/nominal_atoms.ML Tue Jan 25 09:45:45 2011 +0100 @@ -172,26 +172,31 @@ (* overloades then the general swap-function *) val (swap_eqs, thy3) = fold_map (fn (ak_name, T) => fn thy => let + val thy' = Sign.add_path "rec" thy; val swapT = HOLogic.mk_prodT (T, T) --> T --> T; - val swap_name = Sign.full_bname thy ("swap_" ^ ak_name); + val swap_name = "swap_" ^ ak_name; + val full_swap_name = Sign.full_bname thy' swap_name; val a = Free ("a", T); val b = Free ("b", T); val c = Free ("c", T); val ab = Free ("ab", HOLogic.mk_prodT (T, T)) val cif = Const ("HOL.If", HOLogic.boolT --> T --> T --> T); - val cswap_akname = Const (swap_name, swapT); + val cswap_akname = Const (full_swap_name, swapT); val cswap = Const ("Nominal.swap", swapT) - val name = "swap_"^ak_name^"_def"; + val name = swap_name ^ "_def"; val def1 = HOLogic.mk_Trueprop (HOLogic.mk_eq - (cswap_akname $ HOLogic.mk_prod (a,b) $ c, + (Free (swap_name, swapT) $ HOLogic.mk_prod (a,b) $ c, cif $ HOLogic.mk_eq (a,c) $ b $ (cif $ HOLogic.mk_eq (b,c) $ a $ c))) val def2 = Logic.mk_equals (cswap $ ab $ c, cswap_akname $ ab $ c) in - thy |> Sign.add_consts_i [(Binding.name ("swap_" ^ ak_name), swapT, NoSyn)] - |> Global_Theory.add_defs_unchecked true [((Binding.name name, def2),[])] - |> snd - |> OldPrimrec.add_primrec_unchecked_i "" [(("", def1),[])] + thy' |> + Primrec.add_primrec_global + [(Binding.name swap_name, SOME swapT, NoSyn)] + [(Attrib.empty_binding, def1)] ||> + Sign.parent_path ||>> + Global_Theory.add_defs_unchecked true + [((Binding.name name, def2), [])] |>> (snd o fst) end) ak_names_types thy1; (* declares a permutation function for every atom-kind acting *) @@ -201,25 +206,29 @@ (* _prm_ (x#xs) a = swap_ x (perm xs a) *) val (prm_eqs, thy4) = fold_map (fn (ak_name, T) => fn thy => let + val thy' = Sign.add_path "rec" thy; val swapT = HOLogic.mk_prodT (T, T) --> T --> T; - val swap_name = Sign.full_bname thy ("swap_" ^ ak_name) + val swap_name = Sign.full_bname thy' ("swap_" ^ ak_name) val prmT = mk_permT T --> T --> T; val prm_name = ak_name ^ "_prm_" ^ ak_name; - val qu_prm_name = Sign.full_bname thy prm_name; + val prm = Free (prm_name, prmT); val x = Free ("x", HOLogic.mk_prodT (T, T)); val xs = Free ("xs", mk_permT T); val a = Free ("a", T) ; val cnil = Const ("List.list.Nil", mk_permT T); - val def1 = HOLogic.mk_Trueprop (HOLogic.mk_eq (Const (qu_prm_name, prmT) $ cnil $ a, a)); + val def1 = HOLogic.mk_Trueprop (HOLogic.mk_eq (prm $ cnil $ a, a)); val def2 = HOLogic.mk_Trueprop (HOLogic.mk_eq - (Const (qu_prm_name, prmT) $ mk_Cons x xs $ a, - Const (swap_name, swapT) $ x $ (Const (qu_prm_name, prmT) $ xs $ a))); + (prm $ mk_Cons x xs $ a, + Const (swap_name, swapT) $ x $ (prm $ xs $ a))); in - thy |> Sign.add_consts_i [(Binding.name prm_name, mk_permT T --> T --> T, NoSyn)] - |> OldPrimrec.add_primrec_unchecked_i "" [(("", def1), []),(("", def2), [])] + thy' |> + Primrec.add_primrec_global + [(Binding.name prm_name, SOME prmT, NoSyn)] + [(Attrib.empty_binding, def1), (Attrib.empty_binding, def2)] ||> + Sign.parent_path end) ak_names_types thy3; (* defines permutation functions for all combinations of atom-kinds; *) @@ -238,13 +247,15 @@ val pi = Free ("pi", mk_permT T); val a = Free ("a", T'); val cperm = Const ("Nominal.perm", mk_permT T --> T' --> T'); - val cperm_def = Const (Sign.full_bname thy' perm_def_name, mk_permT T --> T' --> T'); + val thy'' = Sign.add_path "rec" thy' + val cperm_def = Const (Sign.full_bname thy'' perm_def_name, mk_permT T --> T' --> T'); + val thy''' = Sign.parent_path thy''; val name = ak_name ^ "_prm_" ^ ak_name' ^ "_def"; val def = Logic.mk_equals (cperm $ pi $ a, if ak_name = ak_name' then cperm_def $ pi $ a else a) in - Global_Theory.add_defs_unchecked true [((Binding.name name, def),[])] thy' + Global_Theory.add_defs_unchecked true [((Binding.name name, def), [])] thy''' end) ak_names_types thy) ak_names_types thy4; (* proves that every atom-kind is an instance of at *) diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Nominal/old_primrec.ML --- a/src/HOL/Nominal/old_primrec.ML Mon Jan 24 22:29:50 2011 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,326 +0,0 @@ -(* Title: HOL/Tools/old_primrec.ML - Author: Norbert Voelker, FernUni Hagen - Author: Stefan Berghofer, TU Muenchen - -Package for defining functions on datatypes by primitive recursion. -*) - -signature OLD_PRIMREC = -sig - val unify_consts: theory -> term list -> term list -> term list * term list - val add_primrec: string -> ((bstring * string) * Attrib.src list) list - -> theory -> thm list * theory - val add_primrec_unchecked: string -> ((bstring * string) * Attrib.src list) list - -> theory -> thm list * theory - val add_primrec_i: string -> ((bstring * term) * attribute list) list - -> theory -> thm list * theory - val add_primrec_unchecked_i: string -> ((bstring * term) * attribute list) list - -> theory -> thm list * theory -end; - -structure OldPrimrec : OLD_PRIMREC = -struct - -open Datatype_Aux; - -exception RecError of string; - -fun primrec_err s = error ("Primrec definition error:\n" ^ s); -fun primrec_eq_err thy s eq = - primrec_err (s ^ "\nin\n" ^ quote (Syntax.string_of_term_global thy eq)); - - -(*the following code ensures that each recursive set always has the - same type in all introduction rules*) -fun unify_consts thy cs intr_ts = - (let - fun varify t (i, ts) = - let val t' = map_types (Logic.incr_tvar (i + 1)) (snd (Type.varify_global [] t)) - in (maxidx_of_term t', t'::ts) end; - val (i, cs') = fold_rev varify cs (~1, []); - val (i', intr_ts') = fold_rev varify intr_ts (i, []); - val rec_consts = fold Term.add_consts cs' []; - val intr_consts = fold Term.add_consts intr_ts' []; - fun unify (cname, cT) = - let val consts = map snd (filter (fn (c, _) => c = cname) intr_consts) - in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end; - val (env, _) = fold unify rec_consts (Vartab.empty, i'); - val subst = Type.legacy_freeze o map_types (Envir.norm_type env) - - in (map subst cs', map subst intr_ts') - end) handle Type.TUNIFY => - (warning "Occurrences of recursive constant have non-unifiable types"; (cs, intr_ts)); - - -(* preprocessing of equations *) - -fun process_eqn thy eq rec_fns = - let - val (lhs, rhs) = - if null (Term.add_vars eq []) then - HOLogic.dest_eq (HOLogic.dest_Trueprop eq) - handle TERM _ => raise RecError "not a proper equation" - else raise RecError "illegal schematic variable(s)"; - - val (recfun, args) = strip_comb lhs; - val fnameT = dest_Const recfun handle TERM _ => - raise RecError "function is not declared as constant in theory"; - - val (ls', rest) = take_prefix is_Free args; - val (middle, rs') = take_suffix is_Free rest; - val rpos = length ls'; - - val (constr, cargs') = if null middle then raise RecError "constructor missing" - else strip_comb (hd middle); - val (cname, T) = dest_Const constr - handle TERM _ => raise RecError "ill-formed constructor"; - val (tname, _) = dest_Type (body_type T) handle TYPE _ => - raise RecError "cannot determine datatype associated with function" - - val (ls, cargs, rs) = - (map dest_Free ls', map dest_Free cargs', map dest_Free rs') - handle TERM _ => raise RecError "illegal argument in pattern"; - val lfrees = ls @ rs @ cargs; - - fun check_vars _ [] = () - | check_vars s vars = raise RecError (s ^ commas_quote (map fst vars)) - in - if length middle > 1 then - raise RecError "more than one non-variable in pattern" - else - (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees); - check_vars "extra variables on rhs: " - (subtract (op =) lfrees (map dest_Free (OldTerm.term_frees rhs))); - case AList.lookup (op =) rec_fns fnameT of - NONE => - (fnameT, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eq))]))::rec_fns - | SOME (_, rpos', eqns) => - if AList.defined (op =) eqns cname then - raise RecError "constructor already occurred as pattern" - else if rpos <> rpos' then - raise RecError "position of recursive argument inconsistent" - else - AList.update (op =) (fnameT, (tname, rpos, (cname, (ls, cargs, rs, rhs, eq))::eqns)) - rec_fns) - end - handle RecError s => primrec_eq_err thy s eq; - -fun process_fun thy descr rec_eqns (i, fnameT as (fname, _)) (fnameTs, fnss) = - let - val (_, (tname, _, constrs)) = List.nth (descr, i); - - (* substitute "fname ls x rs" by "y ls rs" for (x, (_, y)) in subs *) - - fun subst [] t fs = (t, fs) - | subst subs (Abs (a, T, t)) fs = - fs - |> subst subs t - |-> (fn t' => pair (Abs (a, T, t'))) - | subst subs (t as (_ $ _)) fs = - let - val (f, ts) = strip_comb t; - in - if is_Const f andalso member (op =) (map fst rec_eqns) (dest_Const f) then - let - val fnameT' as (fname', _) = dest_Const f; - val (_, rpos, _) = the (AList.lookup (op =) rec_eqns fnameT'); - val ls = take rpos ts; - val rest = drop rpos ts; - val (x', rs) = (hd rest, tl rest) - handle Empty => raise RecError ("not enough arguments\ - \ in recursive application\nof function " ^ quote fname' ^ " on rhs"); - val (x, xs) = strip_comb x' - in case AList.lookup (op =) subs x - of NONE => - fs - |> fold_map (subst subs) ts - |-> (fn ts' => pair (list_comb (f, ts'))) - | SOME (i', y) => - fs - |> fold_map (subst subs) (xs @ ls @ rs) - ||> process_fun thy descr rec_eqns (i', fnameT') - |-> (fn ts' => pair (list_comb (y, ts'))) - end - else - fs - |> fold_map (subst subs) (f :: ts) - |-> (fn (f'::ts') => pair (list_comb (f', ts'))) - end - | subst _ t fs = (t, fs); - - (* translate rec equations into function arguments suitable for rec comb *) - - fun trans eqns (cname, cargs) (fnameTs', fnss', fns) = - (case AList.lookup (op =) eqns cname of - NONE => (warning ("No equation for constructor " ^ quote cname ^ - "\nin definition of function " ^ quote fname); - (fnameTs', fnss', (Const ("HOL.undefined", dummyT))::fns)) - | SOME (ls, cargs', rs, rhs, eq) => - let - val recs = filter (is_rec_type o snd) (cargs' ~~ cargs); - val rargs = map fst recs; - val subs = map (rpair dummyT o fst) - (rev (Term.rename_wrt_term rhs rargs)); - val (rhs', (fnameTs'', fnss'')) = - (subst (map (fn ((x, y), z) => - (Free x, (body_index y, Free z))) - (recs ~~ subs)) rhs (fnameTs', fnss')) - handle RecError s => primrec_eq_err thy s eq - in (fnameTs'', fnss'', - (list_abs_free (cargs' @ subs @ ls @ rs, rhs'))::fns) - end) - - in (case AList.lookup (op =) fnameTs i of - NONE => - if exists (equal fnameT o snd) fnameTs then - raise RecError ("inconsistent functions for datatype " ^ quote tname) - else - let - val (_, _, eqns) = the (AList.lookup (op =) rec_eqns fnameT); - val (fnameTs', fnss', fns) = fold_rev (trans eqns) constrs - ((i, fnameT)::fnameTs, fnss, []) - in - (fnameTs', (i, (fname, #1 (snd (hd eqns)), fns))::fnss') - end - | SOME fnameT' => - if fnameT = fnameT' then (fnameTs, fnss) - else raise RecError ("inconsistent functions for datatype " ^ quote tname)) - end; - - -(* prepare functions needed for definitions *) - -fun get_fns fns ((i : int, (tname, _, constrs)), rec_name) (fs, defs) = - case AList.lookup (op =) fns i of - NONE => - let - val dummy_fns = map (fn (_, cargs) => Const ("HOL.undefined", - replicate (length cargs + length (filter is_rec_type cargs)) - dummyT ---> HOLogic.unitT)) constrs; - val _ = warning ("No function definition for datatype " ^ quote tname) - in - (dummy_fns @ fs, defs) - end - | SOME (fname, ls, fs') => (fs' @ fs, (fname, ls, rec_name, tname) :: defs); - - -(* make definition *) - -fun make_def thy fs (fname, ls, rec_name, tname) = - let - val rhs = fold_rev (fn T => fn t => Abs ("", T, t)) - ((map snd ls) @ [dummyT]) - (list_comb (Const (rec_name, dummyT), - fs @ map Bound (0 ::(length ls downto 1)))) - val def_name = Long_Name.base_name fname ^ "_" ^ Long_Name.base_name tname ^ "_def"; - val def_prop = - singleton (Syntax.check_terms (ProofContext.init_global thy)) - (Logic.mk_equals (Const (fname, dummyT), rhs)); - in (def_name, def_prop) end; - - -(* find datatypes which contain all datatypes in tnames' *) - -fun find_dts (dt_info : info Symtab.table) _ [] = [] - | find_dts dt_info tnames' (tname::tnames) = - (case Symtab.lookup dt_info tname of - NONE => primrec_err (quote tname ^ " is not a datatype") - | SOME dt => - if subset (op =) (tnames', map (#1 o snd) (#descr dt)) then - (tname, dt)::(find_dts dt_info tnames' tnames) - else find_dts dt_info tnames' tnames); - -fun prepare_induct ({descr, induct, ...}: info) rec_eqns = - let - fun constrs_of (_, (_, _, cs)) = - map (fn (cname:string, (_, cargs, _, _, _)) => (cname, map fst cargs)) cs; - val params_of = these o AList.lookup (op =) (maps constrs_of rec_eqns); - in - induct - |> Rule_Cases.rename_params (map params_of (maps (map #1 o #3 o #2) descr)) - |> Rule_Cases.save induct - end; - -local - -fun gen_primrec_i note def alt_name eqns_atts thy = - let - val (eqns, atts) = split_list eqns_atts; - val dt_info = Datatype_Data.get_all thy; - val rec_eqns = fold_rev (process_eqn thy o snd) eqns [] ; - val tnames = distinct (op =) (map (#1 o snd) rec_eqns); - val dts = find_dts dt_info tnames tnames; - val main_fns = - map (fn (tname, {index, ...}) => - (index, - (fst o the o find_first (fn f => (#1 o snd) f = tname)) rec_eqns)) - dts; - val {descr, rec_names, rec_rewrites, ...} = - if null dts then - primrec_err ("datatypes " ^ commas_quote tnames ^ "\nare not mutually recursive") - else snd (hd dts); - val (fnameTs, fnss) = - fold_rev (process_fun thy descr rec_eqns) main_fns ([], []); - val (fs, defs) = fold_rev (get_fns fnss) (descr ~~ rec_names) ([], []); - val defs' = map (make_def thy fs) defs; - val nameTs1 = map snd fnameTs; - val nameTs2 = map fst rec_eqns; - val _ = if eq_set (op =) (nameTs1, nameTs2) then () - else primrec_err ("functions " ^ commas_quote (map fst nameTs2) ^ - "\nare not mutually recursive"); - val primrec_name = - if alt_name = "" then (space_implode "_" (map (Long_Name.base_name o #1) defs)) else alt_name; - val (defs_thms', thy') = - thy - |> Sign.add_path primrec_name - |> fold_map def (map (fn (name, t) => ((name, []), t)) defs'); - val rewrites = (map mk_meta_eq rec_rewrites) @ map snd defs_thms'; - val simps = map (fn (_, t) => Goal.prove_global thy' [] [] t - (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])) eqns; - val (simps', thy'') = - thy' - |> fold_map note ((map fst eqns ~~ atts) ~~ map single simps); - val simps'' = maps snd simps'; - val lhss = map (Logic.varify_global o fst o Logic.dest_equals o snd) defs'; - in - thy'' - |> note (("simps", - [Simplifier.simp_add, Nitpick_Simps.add, Code.add_default_eqn_attribute]), simps'') - |> snd - |> Spec_Rules.add_global Spec_Rules.Equational (lhss, simps) - |> note (("induct", []), [prepare_induct (#2 (hd dts)) rec_eqns]) - |> snd - |> Sign.parent_path - |> pair simps'' - end; - -fun gen_primrec note def alt_name eqns thy = - let - val ((names, strings), srcss) = apfst split_list (split_list eqns); - val atts = map (map (Attrib.attribute thy)) srcss; - val eqn_ts = map (fn s => Syntax.read_prop_global thy s - handle ERROR msg => cat_error msg ("The error(s) above occurred for " ^ s)) strings; - val rec_ts = map (fn eq => head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop eq))) - handle TERM _ => primrec_eq_err thy "not a proper equation" eq) eqn_ts; - val (_, eqn_ts') = unify_consts thy rec_ts eqn_ts - in - gen_primrec_i note def alt_name (names ~~ eqn_ts' ~~ atts) thy - end; - -fun thy_note ((name, atts), thms) = - Global_Theory.add_thmss [((Binding.name name, thms), atts)] #-> (fn [thms] => pair (name, thms)); -fun thy_def false ((name, atts), t) = - Global_Theory.add_defs false [((Binding.name name, t), atts)] #-> (fn [thm] => pair (name, thm)) - | thy_def true ((name, atts), t) = - Global_Theory.add_defs_unchecked false [((Binding.name name, t), atts)] #-> (fn [thm] => pair (name, thm)); - -in - -val add_primrec = gen_primrec thy_note (thy_def false); -val add_primrec_unchecked = gen_primrec thy_note (thy_def true); -val add_primrec_i = gen_primrec_i thy_note (thy_def false); -val add_primrec_unchecked_i = gen_primrec_i thy_note (thy_def true); - -end; - -end; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Power.thy --- a/src/HOL/Power.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Power.thy Tue Jan 25 09:45:45 2011 +0100 @@ -297,7 +297,7 @@ assume "~ a \ b" then have "b < a" by (simp only: linorder_not_le) then have "b ^ Suc n < a ^ Suc n" - by (simp only: prems power_strict_mono) + by (simp only: assms power_strict_mono) from le and this show False by (simp add: linorder_not_less [symmetric]) qed diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Predicate.thy --- a/src/HOL/Predicate.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Predicate.thy Tue Jan 25 09:45:45 2011 +0100 @@ -93,10 +93,10 @@ subsubsection {* Top and bottom elements *} lemma bot1E [no_atp, elim!]: "bot x \ P" - by (simp add: bot_fun_def bot_bool_def) + by (simp add: bot_fun_def) lemma bot2E [elim!]: "bot x y \ P" - by (simp add: bot_fun_def bot_bool_def) + by (simp add: bot_fun_def) lemma bot_empty_eq: "bot = (\x. x \ {})" by (auto simp add: fun_eq_iff) @@ -105,64 +105,64 @@ by (auto simp add: fun_eq_iff) lemma top1I [intro!]: "top x" - by (simp add: top_fun_def top_bool_def) + by (simp add: top_fun_def) lemma top2I [intro!]: "top x y" - by (simp add: top_fun_def top_bool_def) + by (simp add: top_fun_def) subsubsection {* Binary intersection *} lemma inf1I [intro!]: "A x ==> B x ==> inf A B x" - by (simp add: inf_fun_def inf_bool_def) + by (simp add: inf_fun_def) lemma inf2I [intro!]: "A x y ==> B x y ==> inf A B x y" - by (simp add: inf_fun_def inf_bool_def) + by (simp add: inf_fun_def) lemma inf1E [elim!]: "inf A B x ==> (A x ==> B x ==> P) ==> P" - by (simp add: inf_fun_def inf_bool_def) + by (simp add: inf_fun_def) lemma inf2E [elim!]: "inf A B x y ==> (A x y ==> B x y ==> P) ==> P" - by (simp add: inf_fun_def inf_bool_def) + by (simp add: inf_fun_def) lemma inf1D1: "inf A B x ==> A x" - by (simp add: inf_fun_def inf_bool_def) + by (simp add: inf_fun_def) lemma inf2D1: "inf A B x y ==> A x y" - by (simp add: inf_fun_def inf_bool_def) + by (simp add: inf_fun_def) lemma inf1D2: "inf A B x ==> B x" - by (simp add: inf_fun_def inf_bool_def) + by (simp add: inf_fun_def) lemma inf2D2: "inf A B x y ==> B x y" - by (simp add: inf_fun_def inf_bool_def) + by (simp add: inf_fun_def) lemma inf_Int_eq: "inf (\x. x \ R) (\x. x \ S) = (\x. x \ R \ S)" - by (simp add: inf_fun_def inf_bool_def mem_def) + by (simp add: inf_fun_def mem_def) lemma inf_Int_eq2 [pred_set_conv]: "inf (\x y. (x, y) \ R) (\x y. (x, y) \ S) = (\x y. (x, y) \ R \ S)" - by (simp add: inf_fun_def inf_bool_def mem_def) + by (simp add: inf_fun_def mem_def) subsubsection {* Binary union *} lemma sup1I1 [elim?]: "A x \ sup A B x" - by (simp add: sup_fun_def sup_bool_def) + by (simp add: sup_fun_def) lemma sup2I1 [elim?]: "A x y \ sup A B x y" - by (simp add: sup_fun_def sup_bool_def) + by (simp add: sup_fun_def) lemma sup1I2 [elim?]: "B x \ sup A B x" - by (simp add: sup_fun_def sup_bool_def) + by (simp add: sup_fun_def) lemma sup2I2 [elim?]: "B x y \ sup A B x y" - by (simp add: sup_fun_def sup_bool_def) + by (simp add: sup_fun_def) lemma sup1E [elim!]: "sup A B x ==> (A x ==> P) ==> (B x ==> P) ==> P" - by (simp add: sup_fun_def sup_bool_def) iprover + by (simp add: sup_fun_def) iprover lemma sup2E [elim!]: "sup A B x y ==> (A x y ==> P) ==> (B x y ==> P) ==> P" - by (simp add: sup_fun_def sup_bool_def) iprover + by (simp add: sup_fun_def) iprover text {* \medskip Classical introduction rule: no commitment to @{text A} vs @@ -170,16 +170,16 @@ *} lemma sup1CI [intro!]: "(~ B x ==> A x) ==> sup A B x" - by (auto simp add: sup_fun_def sup_bool_def) + by (auto simp add: sup_fun_def) lemma sup2CI [intro!]: "(~ B x y ==> A x y) ==> sup A B x y" - by (auto simp add: sup_fun_def sup_bool_def) + by (auto simp add: sup_fun_def) lemma sup_Un_eq: "sup (\x. x \ R) (\x. x \ S) = (\x. x \ R \ S)" - by (simp add: sup_fun_def sup_bool_def mem_def) + by (simp add: sup_fun_def mem_def) lemma sup_Un_eq2 [pred_set_conv]: "sup (\x y. (x, y) \ R) (\x y. (x, y) \ S) = (\x y. (x, y) \ R \ S)" - by (simp add: sup_fun_def sup_bool_def mem_def) + by (simp add: sup_fun_def mem_def) subsubsection {* Intersections of families *} @@ -257,7 +257,7 @@ lemma pred_comp_rel_comp_eq [pred_set_conv]: "((\x y. (x, y) \ r) OO (\x y. (x, y) \ s)) = (\x y. (x, y) \ r O s)" - by (auto simp add: fun_eq_iff elim: pred_compE) + by (auto simp add: fun_eq_iff) subsubsection {* Converse *} @@ -292,12 +292,10 @@ elim: pred_compE dest: conversepD) lemma converse_meet: "(inf r s)^--1 = inf r^--1 s^--1" - by (simp add: inf_fun_def inf_bool_def) - (iprover intro: conversepI ext dest: conversepD) + by (simp add: inf_fun_def) (iprover intro: conversepI ext dest: conversepD) lemma converse_join: "(sup r s)^--1 = sup r^--1 s^--1" - by (simp add: sup_fun_def sup_bool_def) - (iprover intro: conversepI ext dest: conversepD) + by (simp add: sup_fun_def) (iprover intro: conversepI ext dest: conversepD) lemma conversep_noteq [simp]: "(op ~=)^--1 = op ~=" by (auto simp add: fun_eq_iff) @@ -756,7 +754,7 @@ apply (rule ext) apply (simp add: unit_eq) done - from this prems show ?thesis by blast + from this assms show ?thesis by blast qed lemma unit_pred_cases: diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Prolog/ROOT.ML --- a/src/HOL/Prolog/ROOT.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Prolog/ROOT.ML Tue Jan 25 09:45:45 2011 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/Prolog/ROOT.ML - ID: $Id$ Author: David von Oheimb (based on a lecture on Lambda Prolog by Nadathur) *) diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/RComplete.thy --- a/src/HOL/RComplete.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/RComplete.thy Tue Jan 25 09:45:45 2011 +0100 @@ -517,10 +517,10 @@ apply simp done -lemma natfloor_div_nat: "1 <= x ==> y > 0 ==> - natfloor (x / real y) = natfloor x div y" +lemma natfloor_div_nat: + assumes "1 <= x" and "y > 0" + shows "natfloor (x / real y) = natfloor x div y" proof - - assume "1 <= (x::real)" and "(y::nat) > 0" have "natfloor x = (natfloor x) div y * y + (natfloor x) mod y" by simp then have a: "real(natfloor x) = real ((natfloor x) div y) * real y + @@ -535,8 +535,7 @@ by simp also have "... = real((natfloor x) div y) + real((natfloor x) mod y) / real y + (x - real(natfloor x)) / real y" - by (auto simp add: algebra_simps add_divide_distrib - diff_divide_distrib prems) + by (auto simp add: algebra_simps add_divide_distrib diff_divide_distrib) finally have "natfloor (x / real y) = natfloor(...)" by simp also have "... = natfloor(real((natfloor x) mod y) / real y + (x - real(natfloor x)) / real y + real((natfloor x) div y))" @@ -547,11 +546,11 @@ apply (rule add_nonneg_nonneg) apply (rule divide_nonneg_pos) apply simp - apply (simp add: prems) + apply (simp add: assms) apply (rule divide_nonneg_pos) apply (simp add: algebra_simps) apply (rule real_natfloor_le) - apply (insert prems, auto) + using assms apply auto done also have "natfloor(real((natfloor x) mod y) / real y + (x - real(natfloor x)) / real y) = 0" @@ -560,13 +559,13 @@ apply (rule add_nonneg_nonneg) apply (rule divide_nonneg_pos) apply force - apply (force simp add: prems) + apply (force simp add: assms) apply (rule divide_nonneg_pos) apply (simp add: algebra_simps) apply (rule real_natfloor_le) - apply (auto simp add: prems) - apply (insert prems, arith) - apply (simp add: add_divide_distrib [THEN sym]) + apply (auto simp add: assms) + using assms apply arith + using assms apply (simp add: add_divide_distrib [THEN sym]) apply (subgoal_tac "real y = real y - 1 + 1") apply (erule ssubst) apply (rule add_le_less_mono) diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/RealDef.thy --- a/src/HOL/RealDef.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/RealDef.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1200,7 +1200,7 @@ lemma real_of_int_div_aux: "d ~= 0 ==> (real (x::int)) / (real d) = real (x div d) + (real (x mod d)) / (real d)" proof - - assume "d ~= 0" + assume d: "d ~= 0" have "x = (x div d) * d + x mod d" by auto then have "real x = real (x div d) * real d + real(x mod d)" @@ -1208,7 +1208,7 @@ then have "real x / real d = ... / real d" by simp then show ?thesis - by (auto simp add: add_divide_distrib algebra_simps prems) + by (auto simp add: add_divide_distrib algebra_simps d) qed lemma real_of_int_div: "(d::int) ~= 0 ==> d dvd n ==> @@ -1353,7 +1353,7 @@ lemma real_of_nat_div_aux: "0 < d ==> (real (x::nat)) / (real d) = real (x div d) + (real (x mod d)) / (real d)" proof - - assume "0 < d" + assume d: "0 < d" have "x = (x div d) * d + x mod d" by auto then have "real x = real (x div d) * real d + real(x mod d)" @@ -1361,7 +1361,7 @@ then have "real x / real d = \ / real d" by simp then show ?thesis - by (auto simp add: add_divide_distrib algebra_simps prems) + by (auto simp add: add_divide_distrib algebra_simps d) qed lemma real_of_nat_div: "0 < (d::nat) ==> d dvd n ==> diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SMT.thy --- a/src/HOL/SMT.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/SMT.thy Tue Jan 25 09:45:45 2011 +0100 @@ -185,7 +185,7 @@ @{text yes}. *} -declare [[ smt_solver = cvc3 ]] +declare [[ smt_solver = z3 ]] text {* Since SMT solvers are potentially non-terminating, there is a timeout diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SMT_Examples/SMT_Examples.thy --- a/src/HOL/SMT_Examples/SMT_Examples.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/SMT_Examples/SMT_Examples.thy Tue Jan 25 09:45:45 2011 +0100 @@ -8,7 +8,7 @@ imports Complex_Main begin -declare [[smt_solver=z3, smt_oracle=false]] +declare [[smt_oracle=false]] declare [[smt_certificates="SMT_Examples.certs"]] declare [[smt_fixed=true]] diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SMT_Examples/SMT_Tests.thy --- a/src/HOL/SMT_Examples/SMT_Tests.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/SMT_Examples/SMT_Tests.thy Tue Jan 25 09:45:45 2011 +0100 @@ -8,7 +8,7 @@ imports Complex_Main begin -declare [[smt_solver=z3, smt_oracle=false]] +declare [[smt_oracle=false]] declare [[smt_certificates="SMT_Tests.certs"]] declare [[smt_fixed=true]] diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SMT_Examples/SMT_Word_Examples.thy --- a/src/HOL/SMT_Examples/SMT_Word_Examples.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/SMT_Examples/SMT_Word_Examples.thy Tue Jan 25 09:45:45 2011 +0100 @@ -8,7 +8,7 @@ imports Word begin -declare [[smt_solver=z3, smt_oracle=true]] +declare [[smt_oracle=true]] declare [[smt_certificates="SMT_Word_Examples.certs"]] declare [[smt_fixed=true]] diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/Gcd/Gcd.adb --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/Gcd/Gcd.adb Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,17 @@ +package body Greatest_Common_Divisor +is + + procedure G_C_D(M, N: in Natural; G: out Natural) + is + C, D, R: Integer; + begin + C := M; D := N; + while D /= 0 loop + --# assert C >= 0 and D > 0 and Gcd(C, D) = Gcd(M, N); + R := C rem D; + C := D; D := R; + end loop; + G := C; + end G_C_D; + +end Greatest_Common_Divisor; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/Gcd/Gcd.ads --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/Gcd/Gcd.ads Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,11 @@ +package Greatest_Common_Divisor +is + + --# function Gcd(A, B: Natural) return Natural; + + procedure G_C_D(M, N: in Natural; G: out Natural); + --# derives G from M, N; + --# pre M >= 0 and N > 0; + --# post G = Gcd(M,N); + +end Greatest_Common_Divisor; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/Gcd/Greatest_Common_Divisor.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/Gcd/Greatest_Common_Divisor.thy Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,36 @@ +(* Title: HOL/SPARK/Examples/Gcd/Greatest_Common_Divisor.thy + Author: Stefan Berghofer + Copyright: secunet Security Networks AG +*) + +theory Greatest_Common_Divisor +imports SPARK GCD +begin + +spark_proof_functions + gcd = "gcd :: int \ int \ int" + +spark_open "greatest_common_divisor/g_c_d.siv" + +spark_vc procedure_g_c_d_4 +proof - + from `0 < d` have "0 \ c mod d" by (rule pos_mod_sign) + with `0 \ c` `0 < d` `c - c sdiv d * d \ 0` show ?C1 + by (simp add: sdiv_pos_pos zmod_zdiv_equality') +next + from `0 \ c` `0 < d` `gcd c d = gcd m n` show ?C2 + by (simp add: sdiv_pos_pos zmod_zdiv_equality' gcd_non_0_int) +qed + +spark_vc procedure_g_c_d_11 +proof - + from `0 \ c` `0 < d` `c - c sdiv d * d = 0` + have "d dvd c" + by (auto simp add: sdiv_pos_pos dvd_def mult_ac) + with `0 < d` `gcd c d = gcd m n` show ?C1 + by simp +qed + +spark_end + +end diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/Gcd/greatest_common_divisor/g_c_d.fdl --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/Gcd/greatest_common_divisor/g_c_d.fdl Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,32 @@ + {*******************************************************} + {FDL Declarations} + {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039} + {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.} + {*******************************************************} + + + {DATE : 29-NOV-2010 14:30:10.98} + + {procedure Greatest_Common_Divisor.G_C_D} + + +title procedure g_c_d; + + function round__(real) : integer; + const natural__base__first : integer = pending; + const natural__base__last : integer = pending; + const integer__base__first : integer = pending; + const integer__base__last : integer = pending; + const natural__first : integer = pending; + const natural__last : integer = pending; + const natural__size : integer = pending; + const integer__first : integer = pending; + const integer__last : integer = pending; + const integer__size : integer = pending; + var m : integer; + var n : integer; + var c : integer; + var d : integer; + function gcd(integer, integer) : integer; + +end; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/Gcd/greatest_common_divisor/g_c_d.rls --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/Gcd/greatest_common_divisor/g_c_d.rls Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,27 @@ + /*********************************************************/ + /*Proof Rule Declarations*/ + /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/ + /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/ + /*********************************************************/ + + + /*DATE : 29-NOV-2010 14:30:10.98*/ + + /*procedure Greatest_Common_Divisor.G_C_D*/ + + +rule_family g_c_d_rules: + X requires [X:any] & + X <= Y requires [X:ire, Y:ire] & + X >= Y requires [X:ire, Y:ire]. + +g_c_d_rules(1): integer__size >= 0 may_be_deduced. +g_c_d_rules(2): integer__first may_be_replaced_by -2147483648. +g_c_d_rules(3): integer__last may_be_replaced_by 2147483647. +g_c_d_rules(4): integer__base__first may_be_replaced_by -2147483648. +g_c_d_rules(5): integer__base__last may_be_replaced_by 2147483647. +g_c_d_rules(6): natural__size >= 0 may_be_deduced. +g_c_d_rules(7): natural__first may_be_replaced_by 0. +g_c_d_rules(8): natural__last may_be_replaced_by 2147483647. +g_c_d_rules(9): natural__base__first may_be_replaced_by -2147483648. +g_c_d_rules(10): natural__base__last may_be_replaced_by 2147483647. diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/Gcd/greatest_common_divisor/g_c_d.siv --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/Gcd/greatest_common_divisor/g_c_d.siv Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,117 @@ +***************************************************************************** + Semantic Analysis of SPARK Text + Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 + Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. +***************************************************************************** + + +CREATED 29-NOV-2010, 14:30:10 SIMPLIFIED 29-NOV-2010, 14:30:11 + +SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 +Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. + +procedure Greatest_Common_Divisor.G_C_D + + + + +For path(s) from start to run-time check associated with statement of line 8: + +procedure_g_c_d_1. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 8: + +procedure_g_c_d_2. +*** true . /* all conclusions proved */ + + +For path(s) from start to assertion of line 10: + +procedure_g_c_d_3. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 10 to assertion of line 10: + +procedure_g_c_d_4. +H1: c >= 0 . +H2: d > 0 . +H3: gcd(c, d) = gcd(m, n) . +H4: m >= 0 . +H5: m <= 2147483647 . +H6: n <= 2147483647 . +H7: n > 0 . +H8: c <= 2147483647 . +H9: d <= 2147483647 . +H10: c - c div d * d >= - 2147483648 . +H11: c - c div d * d <= 2147483647 . +H12: c - c div d * d <> 0 . +H13: integer__size >= 0 . +H14: natural__size >= 0 . + -> +C1: c - c div d * d > 0 . +C2: gcd(d, c - c div d * d) = gcd(m, n) . + + +For path(s) from assertion of line 10 to run-time check associated with + statement of line 11: + +procedure_g_c_d_5. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 10 to run-time check associated with + statement of line 12: + +procedure_g_c_d_6. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 10 to run-time check associated with + statement of line 12: + +procedure_g_c_d_7. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 14: + +procedure_g_c_d_8. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 10 to run-time check associated with + statement of line 14: + +procedure_g_c_d_9. +*** true . /* all conclusions proved */ + + +For path(s) from start to finish: + +procedure_g_c_d_10. +*** true . /* contradiction within hypotheses. */ + + + +For path(s) from assertion of line 10 to finish: + +procedure_g_c_d_11. +H1: c >= 0 . +H2: d > 0 . +H3: gcd(c, d) = gcd(m, n) . +H4: m >= 0 . +H5: m <= 2147483647 . +H6: n <= 2147483647 . +H7: n > 0 . +H8: c <= 2147483647 . +H9: d <= 2147483647 . +H10: c - c div d * d = 0 . +H11: integer__size >= 0 . +H12: natural__size >= 0 . + -> +C1: d = gcd(m, n) . + + diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/Liseq/Liseq.adb --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/Liseq/Liseq.adb Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,55 @@ +------------------------------------------------------------------------------- +-- Longest increasing subsequence of an array of integers +------------------------------------------------------------------------------- + +package body Liseq is + + procedure Liseq_length(A: in Vector; L: in out Vector; maxi: out Integer) + is + maxj,i,j,pmax : Integer; + begin + L(0) := 1; + pmax := 0; + maxi := 1; + i := 1; + while i <= L'Last + --# assert + --# (for all i2 in Integer range 0 .. i-1 => + --# (L(i2) = Liseq_ends_at(A, i2))) and + --# L(pmax) = maxi and L(pmax) = Liseq_prfx(A, i) and + --# 0 <= i and i <= L'Last+1 and 0 <= pmax and pmax < i and + --# A'First = 0 and L'First = 0 and A'Last = L'Last and A'Last < Integer'Last; + loop + if A(i) < A(pmax) then + maxj := 0; + j := 0; + while j < i + --# assert + --# (for all i2 in Integer range 0 .. i-1 => + --# (L(i2) = Liseq_ends_at(A, i2))) and + --# L(pmax) = maxi and L(pmax) = Liseq_prfx(A, I) and + --# 0 <= i and i <= L'Last and 0 <= pmax and pmax < i and + --# 0 <= j and j <= i and + --# maxj = Max_ext (A, i, j) and + --# A'First = 0 and L'First = 0 and A'Last = L'Last and A'Last < Integer'Last; + loop + if (A(j) <= A(i) and + maxj < L(j)) then + maxj := L(j); + end if; + j := j+1; + end loop; + L(i) := maxj+1; + if L(i) > maxi then + maxi := maxi+1; + pmax := i; + end if; + else + maxi := maxi+1; + L(i) := maxi; + pmax := i; + end if; + i := i+1; + end loop; + end Liseq_length; +end Liseq; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/Liseq/Liseq.ads --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/Liseq/Liseq.ads Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,18 @@ +------------------------------------------------------------------------------- +-- Longest increasing subsequence of an array of integers +------------------------------------------------------------------------------- + +package Liseq is + + type Vector is array (Integer range <>) of Integer; + + --# function Liseq_prfx(A: Vector; i: Integer) return Integer; + --# function Liseq_ends_at(A: Vector; i: Integer) return Integer; + --# function Max_ext(A: Vector; i, j: Integer) return Integer; + + procedure Liseq_length(A: in Vector; L: in out Vector; maxi: out Integer); + --# derives maxi, L from A, L; + --# pre A'First = 0 and L'First = 0 and A'Last = L'Last and A'Last < Integer'Last; + --# post maxi = Liseq_prfx (A, A'Last+1); + +end Liseq; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/Liseq/Longest_Increasing_Subsequence.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/Liseq/Longest_Increasing_Subsequence.thy Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,665 @@ +(* Title: HOL/SPARK/Examples/Liseq/Longest_Increasing_Subsequence.thy + Author: Stefan Berghofer + Copyright: secunet Security Networks AG +*) + +theory Longest_Increasing_Subsequence +imports SPARK +begin + +text {* +Set of all increasing subsequences in a prefix of an array +*} + +definition iseq :: "(nat \ 'a::linorder) \ nat \ nat set set" where + "iseq xs l = {is. (\i\is. i < l) \ + (\i\is. \j\is. i \ j \ xs i \ xs j)}" + +text {* +Length of longest increasing subsequence in a prefix of an array +*} + +definition liseq :: "(nat \ 'a::linorder) \ nat \ nat" where + "liseq xs i = Max (card ` iseq xs i)" + +text {* +Length of longest increasing subsequence ending at a particular position +*} + +definition liseq' :: "(nat \ 'a::linorder) \ nat \ nat" where + "liseq' xs i = Max (card ` (iseq xs (Suc i) \ {is. Max is = i}))" + +lemma iseq_finite: "finite (iseq xs i)" + apply (simp add: iseq_def) + apply (rule finite_subset [OF _ + finite_Collect_subsets [of "{j. j < i}"]]) + apply auto + done + +lemma iseq_finite': "is \ iseq xs i \ finite is" + by (auto simp add: iseq_def bounded_nat_set_is_finite) + +lemma iseq_singleton: "i < l \ {i} \ iseq xs l" + by (simp add: iseq_def) + +lemma iseq_trivial: "{} \ iseq xs i" + by (simp add: iseq_def) + +lemma iseq_nonempty: "iseq xs i \ {}" + by (auto intro: iseq_trivial) + +lemma liseq'_ge1: "1 \ liseq' xs x" + apply (simp add: liseq'_def) + apply (subgoal_tac "iseq xs (Suc x) \ {is. Max is = x} \ {}") + apply (simp add: Max_ge_iff iseq_finite) + apply (rule_tac x="{x}" in bexI) + apply (auto intro: iseq_singleton) + done + +lemma liseq_expand: + assumes R: "\is. liseq xs i = card is \ is \ iseq xs i \ + (\js. js \ iseq xs i \ card js \ card is) \ P" + shows "P" +proof - + have "Max (card ` iseq xs i) \ card ` iseq xs i" + by (rule Max_in) (simp_all add: iseq_finite iseq_nonempty) + then obtain js where js: "liseq xs i = card js" and "js \ iseq xs i" + by (rule imageE) (simp add: liseq_def) + moreover { + fix js' + assume "js' \ iseq xs i" + then have "card js' \ card js" + by (simp add: js [symmetric] liseq_def iseq_finite iseq_trivial) + } + ultimately show ?thesis by (rule R) +qed + +lemma liseq'_expand: + assumes R: "\is. liseq' xs i = card is \ is \ iseq xs (Suc i) \ + finite is \ Max is = i \ + (\js. js \ iseq xs (Suc i) \ Max js = i \ card js \ card is) \ + is \ {} \ P" + shows "P" +proof - + have "Max (card ` (iseq xs (Suc i) \ {is. Max is = i})) \ + card ` (iseq xs (Suc i) \ {is. Max is = i})" + by (auto simp add: iseq_finite intro!: iseq_singleton Max_in) + then obtain js where js: "liseq' xs i = card js" and "js \ iseq xs (Suc i)" + and "finite js" and "Max js = i" + by (auto simp add: liseq'_def intro: iseq_finite') + moreover { + fix js' + assume "js' \ iseq xs (Suc i)" "Max js' = i" + then have "card js' \ card js" + by (auto simp add: js [symmetric] liseq'_def iseq_finite intro!: iseq_singleton) + } + note max = this + moreover have "card {i} \ card js" + by (rule max) (simp_all add: iseq_singleton) + then have "js \ {}" by auto + ultimately show ?thesis by (rule R) +qed + +lemma liseq'_ge: + "j = card js \ js \ iseq xs (Suc i) \ Max js = i \ + js \ {} \ j \ liseq' xs i" + by (simp add: liseq'_def iseq_finite) + +lemma liseq'_eq: + "j = card js \ js \ iseq xs (Suc i) \ Max js = i \ + js \ {} \ (\js'. js' \ iseq xs (Suc i) \ Max js' = i \ finite js' \ + js' \ {} \ card js' \ card js) \ + j = liseq' xs i" + by (fastsimp simp add: liseq'_def iseq_finite + intro: Max_eqI [symmetric]) + +lemma liseq_ge: + "j = card js \ js \ iseq xs i \ j \ liseq xs i" + by (auto simp add: liseq_def iseq_finite) + +lemma liseq_eq: + "j = card js \ js \ iseq xs i \ + (\js'. js' \ iseq xs i \ finite js' \ + js' \ {} \ card js' \ card js) \ + j = liseq xs i" + by (fastsimp simp add: liseq_def iseq_finite + intro: Max_eqI [symmetric]) + +lemma max_notin: "finite xs \ Max xs < x \ x \ xs" + by (cases "xs = {}") auto + +lemma iseq_insert: + "xs (Max is) \ xs i \ is \ iseq xs i \ + is \ {i} \ iseq xs (Suc i)" + apply (frule iseq_finite') + apply (cases "is = {}") + apply (auto simp add: iseq_def) + apply (rule order_trans [of _ "xs (Max is)"]) + apply auto + apply (thin_tac "\a\is. a < i") + apply (drule_tac x=ia in bspec) + apply assumption + apply (drule_tac x="Max is" in bspec) + apply (auto intro: Max_in) + done + +lemma iseq_diff: "is \ iseq xs (Suc (Max is)) \ + is - {Max is} \ iseq xs (Suc (Max (is - {Max is})))" + apply (frule iseq_finite') + apply (simp add: iseq_def less_Suc_eq_le) + done + +lemma iseq_butlast: + assumes "js \ iseq xs (Suc i)" and "js \ {}" + and "Max js \ i" + shows "js \ iseq xs i" +proof - + from assms have fin: "finite js" + by (simp add: iseq_finite') + with assms have "Max js \ js" + by auto + with assms have "Max js < i" + by (auto simp add: iseq_def) + with fin assms have "\j\js. j < i" + by simp + with assms show ?thesis + by (simp add: iseq_def) +qed + +lemma iseq_mono: "is \ iseq xs i \ i \ j \ is \ iseq xs j" + by (auto simp add: iseq_def) + +lemma diff_nonempty: + assumes "1 < card is" + shows "is - {i} \ {}" +proof - + from assms have fin: "finite is" by (auto intro: card_ge_0_finite) + with assms fin have "card is - 1 \ card (is - {i})" + by (simp add: card_Diff_singleton_if) + with assms have "0 < card (is - {i})" by simp + then show ?thesis by (simp add: card_gt_0_iff) +qed + +lemma Max_diff: + assumes "1 < card is" + shows "Max (is - {Max is}) < Max is" +proof - + from assms have "finite is" by (auto intro: card_ge_0_finite) + moreover from assms have "is - {Max is} \ {}" + by (rule diff_nonempty) + ultimately show ?thesis using assms + apply (auto simp add: not_less) + apply (subgoal_tac "a \ Max is") + apply auto + done +qed + +lemma iseq_nth: "js \ iseq xs l \ 1 < card js \ + xs (Max (js - {Max js})) \ xs (Max js)" + apply (auto simp add: iseq_def) + apply (subgoal_tac "Max (js - {Max js}) \ js") + apply (thin_tac "\i\js. i < l") + apply (drule_tac x="Max (js - {Max js})" in bspec) + apply assumption + apply (drule_tac x="Max js" in bspec) + using card_gt_0_iff [of js] + apply simp + using Max_diff [of js] + apply simp + using Max_in [of "js - {Max js}", OF _ diff_nonempty] card_gt_0_iff [of js] + apply auto + done + +lemma card_leq1_singleton: + assumes "finite xs" "xs \ {}" "card xs \ 1" + obtains x where "xs = {x}" + using assms + by induct simp_all + +lemma longest_iseq1: + "liseq' xs i = + Max ({0} \ {liseq' xs j |j. j < i \ xs j \ xs i}) + 1" +proof - + have "Max ({0} \ {liseq' xs j |j. j < i \ xs j \ xs i}) = liseq' xs i - 1" + proof (rule Max_eqI) + fix y + assume "y \ {0} \ {liseq' xs j |j. j < i \ xs j \ xs i}" + then show "y \ liseq' xs i - 1" + proof + assume "y \ {liseq' xs j |j. j < i \ xs j \ xs i}" + then obtain j where j: "j < i" "xs j \ xs i" "y = liseq' xs j" + by auto + have "liseq' xs j + 1 \ liseq' xs i" + proof (rule liseq'_expand) + fix "is" + assume H: "liseq' xs j = card is" "is \ iseq xs (Suc j)" + "finite is" "Max is = j" "is \ {}" + from H j have "card is + 1 = card (is \ {i})" + by (simp add: card_insert max_notin) + moreover { + from H j have "xs (Max is) \ xs i" by simp + moreover from `j < i` have "Suc j \ i" by simp + with `is \ iseq xs (Suc j)` have "is \ iseq xs i" + by (rule iseq_mono) + ultimately have "is \ {i} \ iseq xs (Suc i)" + by (rule iseq_insert) + } moreover from H j have "Max (is \ {i}) = i" by simp + moreover have "is \ {i} \ {}" by simp + ultimately have "card is + 1 \ liseq' xs i" + by (rule liseq'_ge) + with H show ?thesis by simp + qed + with j show "y \ liseq' xs i - 1" + by simp + qed simp + next + have "liseq' xs i \ 1 \ + (\j. liseq' xs i - 1 = liseq' xs j \ j < i \ xs j \ xs i)" + proof (rule liseq'_expand) + fix "is" + assume H: "liseq' xs i = card is" "is \ iseq xs (Suc i)" + "finite is" "Max is = i" "is \ {}" + assume R: "\js. js \ iseq xs (Suc i) \ Max js = i \ + card js \ card is" + show ?thesis + proof (cases "card is \ 1") + case True with H show ?thesis by simp + next + case False + then have "1 < card is" by simp + then have "Max (is - {Max is}) < Max is" + by (rule Max_diff) + from `is \ iseq xs (Suc i)` `1 < card is` + have "xs (Max (is - {Max is})) \ xs (Max is)" + by (rule iseq_nth) + have "card is - 1 = liseq' xs (Max (is - {i}))" + proof (rule liseq'_eq) + from `Max is = i` [symmetric] `finite is` `is \ {}` + show "card is - 1 = card (is - {i})" by simp + next + from `is \ iseq xs (Suc i)` `Max is = i` [symmetric] + show "is - {i} \ iseq xs (Suc (Max (is - {i})))" + by simp (rule iseq_diff) + next + from `1 < card is` + show "is - {i} \ {}" by (rule diff_nonempty) + next + fix js + assume "js \ iseq xs (Suc (Max (is - {i})))" + "Max js = Max (is - {i})" "finite js" "js \ {}" + from `xs (Max (is - {Max is})) \ xs (Max is)` + `Max js = Max (is - {i})` `Max is = i` + have "xs (Max js) \ xs i" by simp + moreover from `Max is = i` `Max (is - {Max is}) < Max is` + have "Suc (Max (is - {i})) \ i" + by simp + with `js \ iseq xs (Suc (Max (is - {i})))` + have "js \ iseq xs i" + by (rule iseq_mono) + ultimately have "js \ {i} \ iseq xs (Suc i)" + by (rule iseq_insert) + moreover from `js \ {}` `finite js` `Max js = Max (is - {i})` + `Max is = i` [symmetric] `Max (is - {Max is}) < Max is` + have "Max (js \ {i}) = i" + by simp + ultimately have "card (js \ {i}) \ card is" by (rule R) + moreover from `Max is = i` [symmetric] `finite js` + `Max (is - {Max is}) < Max is` `Max js = Max (is - {i})` + have "i \ js" by (simp add: max_notin) + with `finite js` + have "card (js \ {i}) = card ((js \ {i}) - {i}) + 1" + by simp + ultimately show "card js \ card (is - {i})" + using `i \ js` `Max is = i` [symmetric] `is \ {}` `finite is` + by simp + qed simp + with H `Max (is - {Max is}) < Max is` + `xs (Max (is - {Max is})) \ xs (Max is)` + show ?thesis by auto + qed + qed + then show "liseq' xs i - 1 \ {0} \ + {liseq' xs j |j. j < i \ xs j \ xs i}" by simp + qed simp + moreover have "1 \ liseq' xs i" by (rule liseq'_ge1) + ultimately show ?thesis by simp +qed + +lemma longest_iseq2': "liseq xs i < liseq' xs i \ + liseq xs (Suc i) = liseq' xs i" + apply (rule_tac xs=xs and i=i in liseq'_expand) + apply simp + apply (rule liseq_eq [symmetric]) + apply (rule refl) + apply assumption + apply (case_tac "Max js' = i") + apply simp + apply (drule_tac js=js' in iseq_butlast) + apply assumption+ + apply (drule_tac js=js' in liseq_ge [OF refl]) + apply simp + done + +lemma longest_iseq2: "liseq xs i < liseq' xs i \ + liseq xs i + 1 = liseq' xs i" + apply (rule_tac xs=xs and i=i in liseq'_expand) + apply simp + apply (rule_tac xs=xs and i=i in liseq_expand) + apply (drule_tac s="Max is" in sym) + apply simp + apply (case_tac "card is \ 1") + apply simp + apply (drule iseq_diff) + apply (drule_tac i="Suc (Max (is - {Max is}))" and j="Max is" in iseq_mono) + apply (simp add: less_eq_Suc_le [symmetric]) + apply (rule Max_diff) + apply simp + apply (drule_tac x="is - {Max is}" in meta_spec, + drule meta_mp, assumption) + apply simp + done + +lemma longest_iseq3: + "liseq xs j = liseq' xs i \ xs i \ xs j \ i < j \ + liseq xs (Suc j) = liseq xs j + 1" + apply (rule_tac xs=xs and i=j in liseq_expand) + apply simp + apply (rule_tac xs=xs and i=i in liseq'_expand) + apply simp + apply (rule_tac js="isa \ {j}" in liseq_eq [symmetric]) + apply (simp add: card_insert card_Diff_singleton_if max_notin) + apply (rule iseq_insert) + apply simp + apply (erule iseq_mono) + apply simp + apply (case_tac "j = Max js'") + apply simp + apply (drule iseq_diff) + apply (drule_tac x="js' - {j}" in meta_spec) + apply (drule meta_mp) + apply simp + apply (case_tac "card js' \ 1") + apply (erule_tac xs=js' in card_leq1_singleton) + apply assumption+ + apply (simp add: iseq_trivial) + apply (erule iseq_mono) + apply (simp add: less_eq_Suc_le [symmetric]) + apply (rule Max_diff) + apply simp + apply (rule le_diff_iff [THEN iffD1, of 1]) + apply (simp add: card_0_eq [symmetric] del: card_0_eq) + apply (simp add: card_insert) + apply (subgoal_tac "card (js' - {j}) = card js' - 1") + apply (simp add: card_insert card_Diff_singleton_if max_notin) + apply (frule_tac A=js' in Max_in) + apply assumption + apply (simp add: card_Diff_singleton_if) + apply (drule_tac js=js' in iseq_butlast) + apply assumption + apply (erule not_sym) + apply (drule_tac x=js' in meta_spec) + apply (drule meta_mp) + apply assumption + apply (simp add: card_insert_disjoint max_notin) + done + +lemma longest_iseq4: + "liseq xs j = liseq' xs i \ xs i \ xs j \ i < j \ + liseq' xs j = liseq' xs i + 1" + apply (rule_tac xs=xs and i=j in liseq_expand) + apply simp + apply (rule_tac xs=xs and i=i in liseq'_expand) + apply simp + apply (rule_tac js="isa \ {j}" in liseq'_eq [symmetric]) + apply (simp add: card_insert card_Diff_singleton_if max_notin) + apply (rule iseq_insert) + apply simp + apply (erule iseq_mono) + apply simp + apply simp + apply simp + apply (drule_tac s="Max js'" in sym) + apply simp + apply (drule iseq_diff) + apply (drule_tac x="js' - {j}" in meta_spec) + apply (drule meta_mp) + apply simp + apply (case_tac "card js' \ 1") + apply (erule_tac xs=js' in card_leq1_singleton) + apply assumption+ + apply (simp add: iseq_trivial) + apply (erule iseq_mono) + apply (simp add: less_eq_Suc_le [symmetric]) + apply (rule Max_diff) + apply simp + apply (rule le_diff_iff [THEN iffD1, of 1]) + apply (simp add: card_0_eq [symmetric] del: card_0_eq) + apply (simp add: card_insert) + apply (subgoal_tac "card (js' - {j}) = card js' - 1") + apply (simp add: card_insert card_Diff_singleton_if max_notin) + apply (frule_tac A=js' in Max_in) + apply assumption + apply (simp add: card_Diff_singleton_if) + done + +lemma longest_iseq5: "liseq' xs i \ liseq xs i \ + liseq xs (Suc i) = liseq xs i" + apply (rule_tac i=i and xs=xs in liseq'_expand) + apply simp + apply (rule_tac xs=xs and i=i in liseq_expand) + apply simp + apply (rule liseq_eq [symmetric]) + apply (rule refl) + apply (erule iseq_mono) + apply simp + apply (case_tac "Max js' = i") + apply (drule_tac x=js' in meta_spec) + apply simp + apply (drule iseq_butlast, assumption, assumption) + apply simp + done + +lemma liseq_empty: "liseq xs 0 = 0" + apply (rule_tac js="{}" in liseq_eq [symmetric]) + apply simp + apply (rule iseq_trivial) + apply (simp add: iseq_def) + done + +lemma liseq'_singleton: "liseq' xs 0 = 1" + by (simp add: longest_iseq1 [of _ 0]) + +lemma liseq_singleton: "liseq xs (Suc 0) = Suc 0" + by (simp add: longest_iseq2' liseq_empty liseq'_singleton) + +lemma liseq'_Suc_unfold: + "A j \ x \ + (insert 0 {liseq' A j' |j'. j' < Suc j \ A j' \ x}) = + (insert 0 {liseq' A j' |j'. j' < j \ A j' \ x}) \ + {liseq' A j}" + by (auto simp add: less_Suc_eq) + +lemma liseq'_Suc_unfold': + "\ (A j \ x) \ + {liseq' A j' |j'. j' < Suc j \ A j' \ x} = + {liseq' A j' |j'. j' < j \ A j' \ x}" + by (auto simp add: less_Suc_eq) + +lemma iseq_card_limit: + assumes "is \ iseq A i" + shows "card is \ i" +proof - + from assms have "is \ {0.. card {0.. i" + by (rule_tac xs=A and i=i in liseq_expand) + (simp add: iseq_card_limit) + +lemma liseq'_limit: "liseq' A i \ i + 1" + by (rule_tac xs=A and i=i in liseq'_expand) + (simp add: iseq_card_limit) + +definition max_ext :: "(nat \ 'a::linorder) \ nat \ nat \ nat" where + "max_ext A i j = Max ({0} \ {liseq' A j' |j'. j' < j \ A j' \ A i})" + +lemma max_ext_limit: "max_ext A i j \ j" + apply (auto simp add: max_ext_def) + apply (drule Suc_leI) + apply (cut_tac i=j' and A=A in liseq'_limit) + apply simp + done + + +text {* Proof functions *} + +abbreviation (input) + "arr_conv a \ (\n. a (int n))" + +lemma idx_conv_suc: + "0 \ i \ nat (i + 1) = nat i + 1" + by simp + +abbreviation liseq_ends_at' :: "(int \ 'a::linorder) \ int \ int" where + "liseq_ends_at' A i \ int (liseq' (\l. A (int l)) (nat i))" + +abbreviation liseq_prfx' :: "(int \ 'a::linorder) \ int \ int" where + "liseq_prfx' A i \ int (liseq (\l. A (int l)) (nat i))" + +abbreviation max_ext' :: "(int \ 'a::linorder) \ int \ int \ int" where + "max_ext' A i j \ int (max_ext (\l. A (int l)) (nat i) (nat j))" + +spark_proof_functions + liseq_ends_at = "liseq_ends_at' :: (int \ int) \ int \ int" + liseq_prfx = "liseq_prfx' :: (int \ int) \ int \ int" + max_ext = "max_ext' :: (int \ int) \ int \ int \ int" + + +text {* The verification conditions *} + +spark_open "liseq/liseq_length.siv" + +spark_vc procedure_liseq_length_5 + by (simp_all add: liseq_singleton liseq'_singleton) + +spark_vc procedure_liseq_length_6 +proof - + from H1 H2 H3 H4 + have eq: "liseq (arr_conv a) (nat i) = + liseq' (arr_conv a) (nat pmax)" + by simp + from H14 H3 H4 + have pmax1: "arr_conv a (nat pmax) \ arr_conv a (nat i)" + by simp + from H3 H4 have pmax2: "nat pmax < nat i" + by simp + { + fix i2 + assume i2: "0 \ i2" "i2 \ i" + have "(l(i := l pmax + 1)) i2 = + int (liseq' (arr_conv a) (nat i2))" + proof (cases "i2 = i") + case True + from eq pmax1 pmax2 have "liseq' (arr_conv a) (nat i) = + liseq' (arr_conv a) (nat pmax) + 1" + by (rule longest_iseq4) + with True H1 H3 H4 show ?thesis + by simp + next + case False + with H1 i2 show ?thesis + by simp + qed + } + then show ?C1 by simp + from eq pmax1 pmax2 + have "liseq (arr_conv a) (Suc (nat i)) = + liseq (arr_conv a) (nat i) + 1" + by (rule longest_iseq3) + with H2 H3 H4 show ?C2 + by (simp add: idx_conv_suc) +qed + +spark_vc procedure_liseq_length_7 +proof - + from H1 show ?C1 + by (simp add: max_ext_def longest_iseq1 [of _ "nat i"]) + from H6 + have m: "max_ext (arr_conv a) (nat i) (nat i) + 1 = + liseq' (arr_conv a) (nat i)" + by (simp add: max_ext_def longest_iseq1 [of _ "nat i"]) + with H2 H18 + have gt: "liseq (arr_conv a) (nat i) < liseq' (arr_conv a) (nat i)" + by simp + then have "liseq' (arr_conv a) (nat i) = liseq (arr_conv a) (nat i) + 1" + by (rule longest_iseq2 [symmetric]) + with H2 m show ?C2 by simp + from gt have "liseq (arr_conv a) (Suc (nat i)) = liseq' (arr_conv a) (nat i)" + by (rule longest_iseq2') + with m H6 show ?C3 by (simp add: idx_conv_suc) +qed + +spark_vc procedure_liseq_length_8 +proof - + { + fix i2 + assume i2: "0 \ i2" "i2 \ i" + have "(l(i := max_ext' a i i + 1)) i2 = + int (liseq' (arr_conv a) (nat i2))" + proof (cases "i2 = i") + case True + with H1 show ?thesis + by (simp add: max_ext_def longest_iseq1 [of _ "nat i"]) + next + case False + with H1 i2 show ?thesis by simp + qed + } + then show ?C1 by simp + from H2 H6 H18 + have "liseq' (arr_conv a) (nat i) \ liseq (arr_conv a) (nat i)" + by (simp add: max_ext_def longest_iseq1 [of _ "nat i"]) + then have "liseq (arr_conv a) (Suc (nat i)) = liseq (arr_conv a) (nat i)" + by (rule longest_iseq5) + with H2 H6 show ?C2 by (simp add: idx_conv_suc) +qed + +spark_vc procedure_liseq_length_12 + by (simp add: max_ext_def) + +spark_vc procedure_liseq_length_13 + using H1 H6 H13 H21 H22 + by (simp add: max_ext_def + idx_conv_suc liseq'_Suc_unfold max_def del: Max_less_iff) + +spark_vc procedure_liseq_length_14 + using H1 H6 H13 H21 + by (cases "a j \ a i") + (simp_all add: max_ext_def + idx_conv_suc liseq'_Suc_unfold liseq'_Suc_unfold') + +spark_vc procedure_liseq_length_19 + using H3 H4 H5 H8 H9 + apply (rule_tac y="int (nat i)" in order_trans) + apply (cut_tac A="arr_conv a" and i="nat i" and j="nat i" in max_ext_limit) + apply simp_all + done + +spark_vc procedure_liseq_length_23 + using H2 H3 H4 H7 H8 H11 + apply (rule_tac y="int (nat i)" in order_trans) + apply (cut_tac A="arr_conv a" and i="nat i" in liseq_limit) + apply simp_all + done + +spark_vc procedure_liseq_length_29 + using H2 H3 H8 H13 + by (simp add: add1_zle_eq [symmetric]) + +spark_end + +end diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/Liseq/liseq/liseq_length.fdl --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/Liseq/liseq/liseq_length.fdl Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,37 @@ + {*******************************************************} + {FDL Declarations} + {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039} + {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.} + {*******************************************************} + + + {DATE : 29-NOV-2010 14:30:13.02} + + {procedure Liseq.Liseq_length} + + +title procedure liseq_length; + + function round__(real) : integer; + type vector = array [integer] of integer; + const integer__base__first : integer = pending; + const integer__base__last : integer = pending; + const l__index__subtype__1__first : integer = pending; + const l__index__subtype__1__last : integer = pending; + const a__index__subtype__1__first : integer = pending; + const a__index__subtype__1__last : integer = pending; + const integer__first : integer = pending; + const integer__last : integer = pending; + const integer__size : integer = pending; + var a : vector; + var l : vector; + var maxi : integer; + var maxj : integer; + var i : integer; + var j : integer; + var pmax : integer; + function liseq_prfx(vector, integer) : integer; + function liseq_ends_at(vector, integer) : integer; + function max_ext(vector, integer, integer) : integer; + +end; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/Liseq/liseq/liseq_length.rls --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/Liseq/liseq/liseq_length.rls Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,34 @@ + /*********************************************************/ + /*Proof Rule Declarations*/ + /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/ + /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/ + /*********************************************************/ + + + /*DATE : 29-NOV-2010 14:30:13.02*/ + + /*procedure Liseq.Liseq_length*/ + + +rule_family liseq_length_rules: + X requires [X:any] & + X <= Y requires [X:ire, Y:ire] & + X >= Y requires [X:ire, Y:ire]. + +liseq_length_rules(1): integer__size >= 0 may_be_deduced. +liseq_length_rules(2): integer__first may_be_replaced_by -2147483648. +liseq_length_rules(3): integer__last may_be_replaced_by 2147483647. +liseq_length_rules(4): integer__base__first may_be_replaced_by -2147483648. +liseq_length_rules(5): integer__base__last may_be_replaced_by 2147483647. +liseq_length_rules(6): a__index__subtype__1__first >= integer__first may_be_deduced. +liseq_length_rules(7): a__index__subtype__1__last <= integer__last may_be_deduced. +liseq_length_rules(8): a__index__subtype__1__first <= + a__index__subtype__1__last may_be_deduced. +liseq_length_rules(9): a__index__subtype__1__last >= integer__first may_be_deduced. +liseq_length_rules(10): a__index__subtype__1__first <= integer__last may_be_deduced. +liseq_length_rules(11): l__index__subtype__1__first >= integer__first may_be_deduced. +liseq_length_rules(12): l__index__subtype__1__last <= integer__last may_be_deduced. +liseq_length_rules(13): l__index__subtype__1__first <= + l__index__subtype__1__last may_be_deduced. +liseq_length_rules(14): l__index__subtype__1__last >= integer__first may_be_deduced. +liseq_length_rules(15): l__index__subtype__1__first <= integer__last may_be_deduced. diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/Liseq/liseq/liseq_length.siv --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/Liseq/liseq/liseq_length.siv Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,547 @@ +***************************************************************************** + Semantic Analysis of SPARK Text + Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 + Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. +***************************************************************************** + + +CREATED 29-NOV-2010, 14:30:13 SIMPLIFIED 29-NOV-2010, 14:30:13 + +SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 +Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. + +procedure Liseq.Liseq_length + + + + +For path(s) from start to run-time check associated with statement of line 11: + +procedure_liseq_length_1. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 12: + +procedure_liseq_length_2. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 13: + +procedure_liseq_length_3. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 14: + +procedure_liseq_length_4. +*** true . /* all conclusions proved */ + + +For path(s) from start to assertion of line 15: + +procedure_liseq_length_5. +H1: a__index__subtype__1__first = 0 . +H2: l__index__subtype__1__first = 0 . +H3: a__index__subtype__1__last = l__index__subtype__1__last . +H4: a__index__subtype__1__last < 2147483647 . +H5: for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 + <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) + and element(a, [i___1]) <= 2147483647) . +H6: for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 + <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) + and element(l, [i___1]) <= 2147483647) . +H7: 0 <= l__index__subtype__1__last . +H8: integer__size >= 0 . +H9: a__index__subtype__1__first <= a__index__subtype__1__last . +H10: l__index__subtype__1__first <= l__index__subtype__1__last . +H11: a__index__subtype__1__first >= - 2147483648 . +H12: a__index__subtype__1__last >= - 2147483648 . +H13: l__index__subtype__1__first >= - 2147483648 . +H14: l__index__subtype__1__last >= - 2147483648 . +H15: a__index__subtype__1__last <= 2147483647 . +H16: a__index__subtype__1__first <= 2147483647 . +H17: l__index__subtype__1__last <= 2147483647 . +H18: l__index__subtype__1__first <= 2147483647 . + -> +C1: for_all(i2_ : integer, 0 <= i2_ and i2_ <= 0 -> element(update(l, [0], 1) + , [i2_]) = liseq_ends_at(a, i2_)) . +C2: 1 = liseq_prfx(a, 1) . + + +For path(s) from assertion of line 15 to assertion of line 15: + +procedure_liseq_length_6. +H1: for_all(i2_ : integer, 0 <= i2_ and i2_ <= i - 1 -> element(l, [i2_]) = + liseq_ends_at(a, i2_)) . +H2: element(l, [pmax]) = liseq_prfx(a, i) . +H3: 0 <= pmax . +H4: pmax < i . +H5: a__index__subtype__1__first = 0 . +H6: l__index__subtype__1__first = 0 . +H7: a__index__subtype__1__last = l__index__subtype__1__last . +H8: a__index__subtype__1__last < 2147483647 . +H9: for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 + <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) + and element(a, [i___1]) <= 2147483647) . +H10: for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 + <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) + and element(l, [i___1]) <= 2147483647) . +H11: i <= l__index__subtype__1__last . +H12: pmax >= a__index__subtype__1__first . +H13: i <= a__index__subtype__1__last . +H14: element(a, [pmax]) <= element(a, [i]) . +H15: element(l, [pmax]) >= - 2147483648 . +H16: element(l, [pmax]) <= 2147483646 . +H17: i >= l__index__subtype__1__first . +H18: i <= 2147483646 . +H19: integer__size >= 0 . +H20: a__index__subtype__1__first <= a__index__subtype__1__last . +H21: l__index__subtype__1__first <= l__index__subtype__1__last . +H22: a__index__subtype__1__first >= - 2147483648 . +H23: a__index__subtype__1__last >= - 2147483648 . +H24: l__index__subtype__1__first >= - 2147483648 . +H25: l__index__subtype__1__last >= - 2147483648 . +H26: a__index__subtype__1__last <= 2147483647 . +H27: a__index__subtype__1__first <= 2147483647 . +H28: l__index__subtype__1__last <= 2147483647 . +H29: l__index__subtype__1__first <= 2147483647 . + -> +C1: for_all(i2_ : integer, 0 <= i2_ and i2_ <= i -> element(update(l, [i], + element(l, [pmax]) + 1), [i2_]) = liseq_ends_at(a, i2_)) . +C2: element(l, [pmax]) + 1 = liseq_prfx(a, i + 1) . + + +For path(s) from assertion of line 26 to assertion of line 15: + +procedure_liseq_length_7. +H1: for_all(i2_ : integer, 0 <= i2_ and i2_ <= i - 1 -> element(l, [i2_]) = + liseq_ends_at(a, i2_)) . +H2: element(l, [pmax]) = liseq_prfx(a, i) . +H3: i <= l__index__subtype__1__last . +H4: 0 <= pmax . +H5: pmax < i . +H6: 0 <= i . +H7: a__index__subtype__1__first = 0 . +H8: l__index__subtype__1__first = 0 . +H9: a__index__subtype__1__last = l__index__subtype__1__last . +H10: a__index__subtype__1__last < 2147483647 . +H11: for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 + <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) + and element(a, [i___1]) <= 2147483647) . +H12: for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 + <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) + and element(l, [i___1]) <= 2147483647) . +H13: i <= 2147483647 . +H14: max_ext(a, i, i) >= - 2147483648 . +H15: max_ext(a, i, i) <= 2147483646 . +H16: i >= l__index__subtype__1__first . +H17: element(l, [pmax]) >= - 2147483648 . +H18: max_ext(a, i, i) + 1 > element(l, [pmax]) . +H19: element(l, [pmax]) <= 2147483646 . +H20: i <= 2147483646 . +H21: integer__size >= 0 . +H22: a__index__subtype__1__first <= a__index__subtype__1__last . +H23: l__index__subtype__1__first <= l__index__subtype__1__last . +H24: a__index__subtype__1__first >= - 2147483648 . +H25: a__index__subtype__1__last >= - 2147483648 . +H26: l__index__subtype__1__first >= - 2147483648 . +H27: l__index__subtype__1__last >= - 2147483648 . +H28: a__index__subtype__1__last <= 2147483647 . +H29: a__index__subtype__1__first <= 2147483647 . +H30: l__index__subtype__1__last <= 2147483647 . +H31: l__index__subtype__1__first <= 2147483647 . + -> +C1: for_all(i2_ : integer, 0 <= i2_ and i2_ <= i -> element(update(l, [i], + max_ext(a, i, i) + 1), [i2_]) = liseq_ends_at(a, i2_)) . +C2: max_ext(a, i, i) + 1 = element(l, [pmax]) + 1 . +C3: max_ext(a, i, i) + 1 = liseq_prfx(a, i + 1) . + + +procedure_liseq_length_8. +H1: for_all(i2_ : integer, 0 <= i2_ and i2_ <= i - 1 -> element(l, [i2_]) = + liseq_ends_at(a, i2_)) . +H2: element(l, [pmax]) = liseq_prfx(a, i) . +H3: i <= l__index__subtype__1__last . +H4: 0 <= pmax . +H5: pmax < i . +H6: 0 <= i . +H7: a__index__subtype__1__first = 0 . +H8: l__index__subtype__1__first = 0 . +H9: a__index__subtype__1__last = l__index__subtype__1__last . +H10: a__index__subtype__1__last < 2147483647 . +H11: for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 + <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) + and element(a, [i___1]) <= 2147483647) . +H12: for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 + <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) + and element(l, [i___1]) <= 2147483647) . +H13: i <= 2147483647 . +H14: max_ext(a, i, i) >= - 2147483648 . +H15: max_ext(a, i, i) <= 2147483646 . +H16: i >= l__index__subtype__1__first . +H17: element(l, [pmax]) <= 2147483647 . +H18: max_ext(a, i, i) + 1 <= element(l, [pmax]) . +H19: i <= 2147483646 . +H20: integer__size >= 0 . +H21: a__index__subtype__1__first <= a__index__subtype__1__last . +H22: l__index__subtype__1__first <= l__index__subtype__1__last . +H23: a__index__subtype__1__first >= - 2147483648 . +H24: a__index__subtype__1__last >= - 2147483648 . +H25: l__index__subtype__1__first >= - 2147483648 . +H26: l__index__subtype__1__last >= - 2147483648 . +H27: a__index__subtype__1__last <= 2147483647 . +H28: a__index__subtype__1__first <= 2147483647 . +H29: l__index__subtype__1__last <= 2147483647 . +H30: l__index__subtype__1__first <= 2147483647 . + -> +C1: for_all(i2_ : integer, 0 <= i2_ and i2_ <= i -> element(update(l, [i], + max_ext(a, i, i) + 1), [i2_]) = liseq_ends_at(a, i2_)) . +C2: element(l, [pmax]) = liseq_prfx(a, i + 1) . + + +For path(s) from assertion of line 15 to run-time check associated with + statement of line 23: + +procedure_liseq_length_9. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 15 to run-time check associated with + statement of line 24: + +procedure_liseq_length_10. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 15 to run-time check associated with + statement of line 25: + +procedure_liseq_length_11. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 15 to assertion of line 26: + +procedure_liseq_length_12. +H1: for_all(i2_ : integer, 0 <= i2_ and i2_ <= i - 1 -> element(l, [i2_]) = + liseq_ends_at(a, i2_)) . +H2: element(l, [pmax]) = liseq_prfx(a, i) . +H3: 0 <= pmax . +H4: pmax < i . +H5: a__index__subtype__1__first = 0 . +H6: l__index__subtype__1__first = 0 . +H7: a__index__subtype__1__last = l__index__subtype__1__last . +H8: a__index__subtype__1__last < 2147483647 . +H9: for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 + <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) + and element(a, [i___1]) <= 2147483647) . +H10: for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 + <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) + and element(l, [i___1]) <= 2147483647) . +H11: i <= l__index__subtype__1__last . +H12: pmax <= 2147483647 . +H13: pmax >= a__index__subtype__1__first . +H14: i <= a__index__subtype__1__last . +H15: element(a, [i]) < element(a, [pmax]) . +H16: integer__size >= 0 . +H17: a__index__subtype__1__first <= a__index__subtype__1__last . +H18: l__index__subtype__1__first <= l__index__subtype__1__last . +H19: a__index__subtype__1__first >= - 2147483648 . +H20: a__index__subtype__1__last >= - 2147483648 . +H21: l__index__subtype__1__first >= - 2147483648 . +H22: l__index__subtype__1__last >= - 2147483648 . +H23: a__index__subtype__1__last <= 2147483647 . +H24: a__index__subtype__1__first <= 2147483647 . +H25: l__index__subtype__1__last <= 2147483647 . +H26: l__index__subtype__1__first <= 2147483647 . + -> +C1: 0 = max_ext(a, i, 0) . + + +For path(s) from assertion of line 26 to assertion of line 26: + +procedure_liseq_length_13. +H1: for_all(i2_ : integer, 0 <= i2_ and i2_ <= i - 1 -> element(l, [i2_]) = + liseq_ends_at(a, i2_)) . +H2: element(l, [pmax]) = liseq_prfx(a, i) . +H3: i <= l__index__subtype__1__last . +H4: 0 <= pmax . +H5: pmax < i . +H6: 0 <= j . +H7: a__index__subtype__1__first = 0 . +H8: l__index__subtype__1__first = 0 . +H9: a__index__subtype__1__last = l__index__subtype__1__last . +H10: a__index__subtype__1__last < 2147483647 . +H11: for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 + <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) + and element(a, [i___1]) <= 2147483647) . +H12: for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 + <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) + and element(l, [i___1]) <= 2147483647) . +H13: j < i . +H14: max_ext(a, i, j) >= - 2147483648 . +H15: max_ext(a, i, j) <= 2147483647 . +H16: j >= l__index__subtype__1__first . +H17: i >= a__index__subtype__1__first . +H18: i <= a__index__subtype__1__last . +H19: j >= a__index__subtype__1__first . +H20: j <= a__index__subtype__1__last . +H21: element(a, [j]) <= element(a, [i]) . +H22: max_ext(a, i, j) < element(l, [j]) . +H23: element(l, [j]) >= - 2147483648 . +H24: element(l, [j]) <= 2147483647 . +H25: j <= 2147483646 . +H26: integer__size >= 0 . +H27: a__index__subtype__1__first <= a__index__subtype__1__last . +H28: l__index__subtype__1__first <= l__index__subtype__1__last . +H29: a__index__subtype__1__first >= - 2147483648 . +H30: a__index__subtype__1__last >= - 2147483648 . +H31: l__index__subtype__1__first >= - 2147483648 . +H32: l__index__subtype__1__last >= - 2147483648 . +H33: a__index__subtype__1__last <= 2147483647 . +H34: a__index__subtype__1__first <= 2147483647 . +H35: l__index__subtype__1__last <= 2147483647 . +H36: l__index__subtype__1__first <= 2147483647 . + -> +C1: element(l, [j]) = max_ext(a, i, j + 1) . + + +procedure_liseq_length_14. +H1: for_all(i2_ : integer, 0 <= i2_ and i2_ <= i - 1 -> element(l, [i2_]) = + liseq_ends_at(a, i2_)) . +H2: element(l, [pmax]) = liseq_prfx(a, i) . +H3: i <= l__index__subtype__1__last . +H4: 0 <= pmax . +H5: pmax < i . +H6: 0 <= j . +H7: a__index__subtype__1__first = 0 . +H8: l__index__subtype__1__first = 0 . +H9: a__index__subtype__1__last = l__index__subtype__1__last . +H10: a__index__subtype__1__last < 2147483647 . +H11: for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 + <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) + and element(a, [i___1]) <= 2147483647) . +H12: for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 + <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) + and element(l, [i___1]) <= 2147483647) . +H13: j < i . +H14: max_ext(a, i, j) >= - 2147483648 . +H15: max_ext(a, i, j) <= 2147483647 . +H16: j >= l__index__subtype__1__first . +H17: i >= a__index__subtype__1__first . +H18: i <= a__index__subtype__1__last . +H19: j >= a__index__subtype__1__first . +H20: j <= a__index__subtype__1__last . +H21: element(a, [i]) < element(a, [j]) or element(l, [j]) <= max_ext(a, i, j) + . +H22: j <= 2147483646 . +H23: integer__size >= 0 . +H24: a__index__subtype__1__first <= a__index__subtype__1__last . +H25: l__index__subtype__1__first <= l__index__subtype__1__last . +H26: a__index__subtype__1__first >= - 2147483648 . +H27: a__index__subtype__1__last >= - 2147483648 . +H28: l__index__subtype__1__first >= - 2147483648 . +H29: l__index__subtype__1__last >= - 2147483648 . +H30: a__index__subtype__1__last <= 2147483647 . +H31: a__index__subtype__1__first <= 2147483647 . +H32: l__index__subtype__1__last <= 2147483647 . +H33: l__index__subtype__1__first <= 2147483647 . + -> +C1: max_ext(a, i, j) = max_ext(a, i, j + 1) . + + +For path(s) from assertion of line 26 to run-time check associated with + statement of line 36: + +procedure_liseq_length_15. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 26 to run-time check associated with + statement of line 38: + +procedure_liseq_length_16. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 26 to run-time check associated with + statement of line 40: + +procedure_liseq_length_17. +*** true . /* all conclusions proved */ + + +procedure_liseq_length_18. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 26 to run-time check associated with + statement of line 42: + +procedure_liseq_length_19. +H1: for_all(i2_ : integer, 0 <= i2_ and i2_ <= i - 1 -> element(l, [i2_]) = + liseq_ends_at(a, i2_)) . +H2: element(l, [pmax]) = liseq_prfx(a, i) . +H3: i <= l__index__subtype__1__last . +H4: 0 <= pmax . +H5: pmax < i . +H6: a__index__subtype__1__first = 0 . +H7: l__index__subtype__1__first = 0 . +H8: a__index__subtype__1__last = l__index__subtype__1__last . +H9: a__index__subtype__1__last < 2147483647 . +H10: for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 + <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) + and element(a, [i___1]) <= 2147483647) . +H11: for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 + <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) + and element(l, [i___1]) <= 2147483647) . +H12: i <= 2147483647 . +H13: max_ext(a, i, i) >= - 2147483648 . +H14: max_ext(a, i, i) <= 2147483647 . +H15: integer__size >= 0 . +H16: a__index__subtype__1__first <= a__index__subtype__1__last . +H17: l__index__subtype__1__first <= l__index__subtype__1__last . +H18: a__index__subtype__1__first >= - 2147483648 . +H19: a__index__subtype__1__last >= - 2147483648 . +H20: l__index__subtype__1__first >= - 2147483648 . +H21: l__index__subtype__1__last >= - 2147483648 . +H22: a__index__subtype__1__last <= 2147483647 . +H23: a__index__subtype__1__first <= 2147483647 . +H24: l__index__subtype__1__last <= 2147483647 . +H25: l__index__subtype__1__first <= 2147483647 . + -> +C1: max_ext(a, i, i) <= 2147483646 . + + +For path(s) from assertion of line 26 to run-time check associated with + statement of line 43: + +procedure_liseq_length_20. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 26 to run-time check associated with + statement of line 44: + +procedure_liseq_length_21. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 26 to run-time check associated with + statement of line 45: + +procedure_liseq_length_22. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 15 to run-time check associated with + statement of line 48: + +procedure_liseq_length_23. +H1: for_all(i2_ : integer, 0 <= i2_ and i2_ <= i - 1 -> element(l, [i2_]) = + liseq_ends_at(a, i2_)) . +H2: element(l, [pmax]) = liseq_prfx(a, i) . +H3: 0 <= pmax . +H4: pmax < i . +H5: a__index__subtype__1__first = 0 . +H6: l__index__subtype__1__first = 0 . +H7: a__index__subtype__1__last = l__index__subtype__1__last . +H8: a__index__subtype__1__last < 2147483647 . +H9: for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 + <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) + and element(a, [i___1]) <= 2147483647) . +H10: for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 + <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) + and element(l, [i___1]) <= 2147483647) . +H11: i <= l__index__subtype__1__last . +H12: pmax <= 2147483647 . +H13: pmax >= a__index__subtype__1__first . +H14: i <= a__index__subtype__1__last . +H15: element(a, [pmax]) <= element(a, [i]) . +H16: element(l, [pmax]) >= - 2147483648 . +H17: element(l, [pmax]) <= 2147483647 . +H18: integer__size >= 0 . +H19: a__index__subtype__1__first <= a__index__subtype__1__last . +H20: l__index__subtype__1__first <= l__index__subtype__1__last . +H21: a__index__subtype__1__first >= - 2147483648 . +H22: a__index__subtype__1__last >= - 2147483648 . +H23: l__index__subtype__1__first >= - 2147483648 . +H24: l__index__subtype__1__last >= - 2147483648 . +H25: a__index__subtype__1__last <= 2147483647 . +H26: a__index__subtype__1__first <= 2147483647 . +H27: l__index__subtype__1__last <= 2147483647 . +H28: l__index__subtype__1__first <= 2147483647 . + -> +C1: element(l, [pmax]) <= 2147483646 . + + +For path(s) from assertion of line 15 to run-time check associated with + statement of line 49: + +procedure_liseq_length_24. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 15 to run-time check associated with + statement of line 50: + +procedure_liseq_length_25. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 15 to run-time check associated with + statement of line 52: + +procedure_liseq_length_26. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 26 to run-time check associated with + statement of line 52: + +procedure_liseq_length_27. +*** true . /* all conclusions proved */ + + +procedure_liseq_length_28. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 15 to finish: + +procedure_liseq_length_29. +H1: for_all(i2_ : integer, 0 <= i2_ and i2_ <= i - 1 -> element(l, [i2_]) = + liseq_ends_at(a, i2_)) . +H2: element(l, [pmax]) = liseq_prfx(a, i) . +H3: i <= l__index__subtype__1__last + 1 . +H4: 0 <= pmax . +H5: pmax < i . +H6: a__index__subtype__1__first = 0 . +H7: l__index__subtype__1__first = 0 . +H8: a__index__subtype__1__last = l__index__subtype__1__last . +H9: a__index__subtype__1__last < 2147483647 . +H10: for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 + <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) + and element(a, [i___1]) <= 2147483647) . +H11: for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 + <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) + and element(l, [i___1]) <= 2147483647) . +H12: i <= 2147483647 . +H13: l__index__subtype__1__last < i . +H14: integer__size >= 0 . +H15: a__index__subtype__1__first <= a__index__subtype__1__last . +H16: l__index__subtype__1__first <= l__index__subtype__1__last . +H17: a__index__subtype__1__first >= - 2147483648 . +H18: a__index__subtype__1__last >= - 2147483648 . +H19: l__index__subtype__1__first >= - 2147483648 . +H20: l__index__subtype__1__last >= - 2147483648 . +H21: a__index__subtype__1__last <= 2147483647 . +H22: a__index__subtype__1__first <= 2147483647 . +H23: l__index__subtype__1__last <= 2147483647 . +H24: l__index__subtype__1__first <= 2147483647 . + -> +C1: element(l, [pmax]) = liseq_prfx(a, a__index__subtype__1__last + 1) . + + diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/README Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,3 @@ +The copyright notice contained in the *.siv, *.fdl, and *.rls files in +the example subdirectories refers to the tools that have been used to +generate the files, but not to the files themselves. diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/F.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/F.thy Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,66 @@ +(* Title: HOL/SPARK/Examples/RIPEMD-160/F.thy + Author: Fabian Immler, TU Muenchen + +Verification of the RIPEMD-160 hash function +*) + +theory F +imports RMD_Specification +begin + +spark_open "rmd/f.siv" + +spark_vc function_f_2 + using assms by simp_all + +spark_vc function_f_3 + using assms by simp_all + +spark_vc function_f_4 + using assms by simp_all + +spark_vc function_f_5 + using assms by simp_all + +spark_vc function_f_6 +proof - + from H8 have "nat j <= 15" by simp + with assms show ?thesis + by (simp add: f_def bwsimps int_word_uint int_mod_eq') +qed + +spark_vc function_f_7 +proof - + from H7 have "16 <= nat j" by simp + moreover from H8 have "nat j <= 31" by simp + ultimately show ?thesis using assms + by (simp add: f_def bwsimps int_word_uint int_mod_eq') +qed + +spark_vc function_f_8 +proof - + from H7 have "32 <= nat j" by simp + moreover from H8 have "nat j <= 47" by simp + ultimately show ?thesis using assms + by (simp add: f_def bwsimps int_word_uint int_mod_eq') +qed + +spark_vc function_f_9 +proof - + from H7 have "48 <= nat j" by simp + moreover from H8 have "nat j <= 63" by simp + ultimately show ?thesis using assms + by (simp add: f_def bwsimps int_word_uint int_mod_eq') +qed + +spark_vc function_f_10 +proof - + from H2 have "nat j <= 79" by simp + moreover from H12 have "64 <= nat j" by simp + ultimately show ?thesis using assms + by (simp add: f_def bwsimps int_word_uint int_mod_eq') +qed + +spark_end + +end diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/Hash.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/Hash.thy Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,100 @@ +(* Title: HOL/SPARK/Examples/RIPEMD-160/Hash.thy + Author: Fabian Immler, TU Muenchen + +Verification of the RIPEMD-160 hash function +*) + +theory Hash +imports RMD_Specification +begin + +spark_open "rmd/hash.siv" + +abbreviation from_chain :: "chain \ RMD.chain" where + "from_chain c \ ( + word_of_int (h0 c), + word_of_int (h1 c), + word_of_int (h2 c), + word_of_int (h3 c), + word_of_int (h4 c))" + +abbreviation to_chain :: "RMD.chain \ chain" where + "to_chain c \ + (let (h0, h1, h2, h3, h4) = c in + (|h0 = uint h0, + h1 = uint h1, + h2 = uint h2, + h3 = uint h3, + h4 = uint h4|))" + +abbreviation round' :: "chain \ block \ chain" where + "round' c b == to_chain (round (\n. word_of_int (b (int n))) (from_chain c))" + +abbreviation rounds' :: "chain \ int \ message \ chain" where + "rounds' h i X == + to_chain (rounds + (\n. \m. word_of_int (X (int n) (int m))) + (from_chain h) + (nat i))" + +abbreviation rmd_hash :: "message \ int \ chain" where + "rmd_hash X i == to_chain (rmd + (\n. \m. word_of_int (X (int n) (int m))) + (nat i))" + +spark_proof_functions + round_spec = round' + rounds = rounds' + rmd_hash = rmd_hash + +spark_vc function_hash_12 + using H1 H6 + by (simp add: + rounds_def rmd_body_def round_def + h_0_def h0_0_def h1_0_def h2_0_def h3_0_def h4_0_def) + + +lemma rounds_step: + assumes "0 <= i" + shows "rounds X b (Suc i) = round (X i) (rounds X b i)" + by (simp add: rounds_def rmd_body_def) + +lemma from_to_id: "from_chain (to_chain C) = C" +proof (cases C) + fix a b c d e f::word32 + assume "C = (a, b, c, d, e)" + thus ?thesis by (cases a) simp +qed + +lemma steps_to_steps': + "round X (foldl a b c) = round X (from_chain (to_chain (foldl a b c)))" + unfolding from_to_id .. + +lemma rounds'_step: + assumes "0 <= i" + shows "rounds' c (i + 1) x = round' (rounds' c i x) (x i)" +proof - + have makesuc: "nat (i + 1) = Suc (nat i)" using assms by simp + show ?thesis using assms + by (simp add: makesuc rounds_def rmd_body_def steps_to_steps') +qed + + +spark_vc function_hash_13 +proof - + have loop_suc: "loop__1__i + 2 = (loop__1__i + 1) + 1" by simp + have "0 <= loop__1__i + 1" using `0 <= loop__1__i` by simp + show ?thesis + unfolding loop_suc + unfolding rounds'_step[OF `0 <= loop__1__i + 1`] + unfolding H1[symmetric] + unfolding H18 .. +qed + + +spark_vc function_hash_17 + unfolding rmd_def H1 rounds_def .. + +spark_end + +end diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/K_L.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/K_L.thy Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,46 @@ +(* Title: HOL/SPARK/Examples/RIPEMD-160/K_L.thy + Author: Fabian Immler, TU Muenchen + +Verification of the RIPEMD-160 hash function +*) + +theory K_L +imports RMD_Specification +begin + +spark_open "rmd/k_l.siv" + +spark_vc function_k_l_6 + using assms by (simp add: K_def) + +spark_vc function_k_l_7 +proof - + from H1 have "16 <= nat j" by simp + moreover from H2 have "nat j <= 31" by simp + ultimately show ?thesis by (simp add: K_def) +qed + +spark_vc function_k_l_8 +proof - + from H1 have "32 <= nat j" by simp + moreover from H2 have "nat j <= 47" by simp + ultimately show ?thesis by (simp add: K_def) +qed + +spark_vc function_k_l_9 +proof - + from H1 have "48 <= nat j" by simp + moreover from H2 have "nat j <= 63" by simp + ultimately show ?thesis by (simp add: K_def) +qed + +spark_vc function_k_l_10 +proof - + from H6 have "64 <= nat j" by simp + moreover from H2 have "nat j <= 79" by simp + ultimately show ?thesis by (simp add: K_def) +qed + +spark_end + +end diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/K_R.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/K_R.thy Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,46 @@ +(* Title: HOL/SPARK/Examples/RIPEMD-160/K_R.thy + Author: Fabian Immler, TU Muenchen + +Verification of the RIPEMD-160 hash function +*) + +theory K_R +imports RMD_Specification +begin + +spark_open "rmd/k_r.siv" + +spark_vc function_k_r_6 + using assms by (simp add: K'_def) + +spark_vc function_k_r_7 +proof- + from H1 have "16 <= nat j" by simp + moreover from H2 have "nat j <= 31" by simp + ultimately show ?thesis by (simp add: K'_def) +qed + +spark_vc function_k_r_8 +proof - + from H1 have "32 <= nat j" by simp + moreover from H2 have "nat j <= 47" by simp + ultimately show ?thesis by (simp add: K'_def) +qed + +spark_vc function_k_r_9 +proof - + from H1 have "48 <= nat j" by simp + moreover from H2 have "nat j <= 63" by simp + ultimately show ?thesis by (simp add: K'_def) +qed + +spark_vc function_k_r_10 +proof - + from H6 have "64 <= nat j" by simp + moreover from H2 have "nat j <= 79" by simp + ultimately show ?thesis by (simp add: K'_def) +qed + +spark_end + +end diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/RMD.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/RMD.thy Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,180 @@ +(* Title: HOL/SPARK/Examples/RIPEMD-160/RMD.thy + Author: Fabian Immler, TU Muenchen + +Verification of the RIPEMD-160 hash function +*) + +theory RMD +imports Word +begin + + +(* all operations are defined on 32-bit words *) + +type_synonym word32 = "32 word" +type_synonym byte = "8 word" +type_synonym perm = "nat \ nat" +type_synonym chain = "word32 * word32 * word32 * word32 * word32" +type_synonym block = "nat \ word32" +type_synonym message = "nat \ block" + +(* nonlinear functions at bit level *) + +definition f::"[nat, word32, word32, word32] => word32" +where +"f j x y z = + (if ( 0 <= j & j <= 15) then x XOR y XOR z + else if (16 <= j & j <= 31) then (x AND y) OR (NOT x AND z) + else if (32 <= j & j <= 47) then (x OR NOT y) XOR z + else if (48 <= j & j <= 63) then (x AND z) OR (y AND NOT z) + else if (64 <= j & j <= 79) then x XOR (y OR NOT z) + else 0)" + + +(* added constants (hexadecimal) *) + +definition K::"nat => word32" +where +"K j = + (if ( 0 <= j & j <= 15) then 0x00000000 + else if (16 <= j & j <= 31) then 0x5A827999 + else if (32 <= j & j <= 47) then 0x6ED9EBA1 + else if (48 <= j & j <= 63) then 0x8F1BBCDC + else if (64 <= j & j <= 79) then 0xA953FD4E + else 0)" + +definition K'::"nat => word32" +where +"K' j = + (if ( 0 <= j & j <= 15) then 0x50A28BE6 + else if (16 <= j & j <= 31) then 0x5C4DD124 + else if (32 <= j & j <= 47) then 0x6D703EF3 + else if (48 <= j & j <= 63) then 0x7A6D76E9 + else if (64 <= j & j <= 79) then 0x00000000 + else 0)" + + +(* selection of message word *) + +definition r_list :: "nat list" + where "r_list = [ + 0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, + 7, 4, 13, 1, 10, 6, 15, 3, 12, 0, 9, 5, 2, 14, 11, 8, + 3, 10, 14, 4, 9, 15, 8, 1, 2, 7, 0, 6, 13, 11, 5, 12, + 1, 9, 11, 10, 0, 8, 12, 4, 13, 3, 7, 15, 14, 5, 6, 2, + 4, 0, 5, 9, 7, 12, 2, 10, 14, 1, 3, 8, 11, 6, 15, 13 + ]" + +definition r'_list :: "nat list" + where "r'_list = [ + 5, 14, 7, 0, 9, 2, 11, 4, 13, 6, 15, 8, 1, 10, 3, 12, + 6, 11, 3, 7, 0, 13, 5, 10, 14, 15, 8, 12, 4, 9, 1, 2, + 15, 5, 1, 3, 7, 14, 6, 9, 11, 8, 12, 2, 10, 0, 4, 13, + 8, 6, 4, 1, 3, 11, 15, 0, 5, 12, 2, 13, 9, 7, 10, 14, + 12, 15, 10, 4, 1, 5, 8, 7, 6, 2, 13, 14, 0, 3, 9, 11 + ]" + +definition r :: perm + where "r j = r_list ! j" + +definition r' :: perm + where "r' j = r'_list ! j" + + +(* amount for rotate left (rol) *) + +definition s_list :: "nat list" + where "s_list = [ + 11, 14, 15, 12, 5, 8, 7, 9, 11, 13, 14, 15, 6, 7, 9, 8, + 7, 6, 8, 13, 11, 9, 7, 15, 7, 12, 15, 9, 11, 7, 13, 12, + 11, 13, 6, 7, 14, 9, 13, 15, 14, 8, 13, 6, 5, 12, 7, 5, + 11, 12, 14, 15, 14, 15, 9, 8, 9, 14, 5, 6, 8, 6, 5, 12, + 9, 15, 5, 11, 6, 8, 13, 12, 5, 12, 13, 14, 11, 8, 5, 6 + ]" + +definition s'_list :: "nat list" + where "s'_list = [ + 8, 9, 9, 11, 13, 15, 15, 5, 7, 7, 8, 11, 14, 14, 12, 6, + 9, 13, 15, 7, 12, 8, 9, 11, 7, 7, 12, 7, 6, 15, 13, 11, + 9, 7, 15, 11, 8, 6, 6, 14, 12, 13, 5, 14, 13, 13, 7, 5, + 15, 5, 8, 11, 14, 14, 6, 14, 6, 9, 12, 9, 12, 5, 15, 8, + 8, 5, 12, 9, 12, 5, 14, 6, 8, 13, 6, 5, 15, 13, 11, 11 + ]" + +definition s :: perm + where "s j = s_list ! j" + +definition s' :: perm + where "s' j = s'_list ! j" + + +(* Initial value (hexadecimal *) + +definition h0_0::word32 where "h0_0 = 0x67452301" +definition h1_0::word32 where "h1_0 = 0xEFCDAB89" +definition h2_0::word32 where "h2_0 = 0x98BADCFE" +definition h3_0::word32 where "h3_0 = 0x10325476" +definition h4_0::word32 where "h4_0 = 0xC3D2E1F0" +definition h_0::chain where "h_0 = (h0_0, h1_0, h2_0, h3_0, h4_0)" + + +definition step_l :: + "[ block, + chain, + nat + ] => chain" + where + "step_l X c j = + (let (A, B, C, D, E) = c in + ((* A *) E, + (* B *) word_rotl (s j) (A + f j B C D + X (r j) + K j) + E, + (* C *) B, + (* D *) word_rotl 10 C, + (* E *) D))" + +definition step_r :: + "[ block, + chain, + nat + ] \ chain" +where + "step_r X c' j = + (let (A', B', C', D', E') = c' in + ((* A' *) E', + (* B' *) word_rotl (s' j) (A' + f (79 - j) B' C' D' + X (r' j) + K' j) + E', + (* C' *) B', + (* D' *) word_rotl 10 C', + (* E' *) D'))" + +definition step_both :: + "[ block, chain * chain, nat ] \ chain * chain" + where + "step_both X cc j = (case cc of (c, c') \ + (step_l X c j, step_r X c' j))" + +definition steps::"[ block, chain * chain, nat] \ chain * chain" + where "steps X cc i = foldl (step_both X) cc [0.. chain" + where "round X h = + (let (h0, h1, h2, h3, h4) = h in + let ((A, B, C, D, E), (A', B', C', D', E')) = steps X (h, h) 80 in + ((* h0 *) h1 + C + D', + (* h1 *) h2 + D + E', + (* h2 *) h3 + E + A', + (* h3 *) h4 + A + B', + (* h4 *) h0 + B + C'))" + +definition rmd_body::"[ message, chain, nat ] => chain" +where + "rmd_body X h i = round (X i) h" + +definition rounds::"message \ chain \ nat \ chain" +where + "rounds X h i = foldl (rmd_body X) h_0 [0.. nat \ chain" +where + "rmd X len = rounds X h_0 len" + +end \ No newline at end of file diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/RMD_Lemmas.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/RMD_Lemmas.thy Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,27 @@ +(* Title: HOL/SPARK/Examples/RIPEMD-160/RMD_Lemmas.thy + Author: Fabian Immler, TU Muenchen + +Verification of the RIPEMD-160 hash function +*) + +theory RMD_Lemmas +imports Main +begin + +definition "fun_of_list i xs g j = + (if j < i \ i + int (length xs) \ j then g j else xs ! nat (j - i))" + +lemma fun_of_list_Nil [simp]: "fun_of_list i [] g = g" + by (auto simp add: fun_eq_iff fun_of_list_def) + +lemma fun_of_list_Cons [simp]: + "fun_of_list i (x # xs) g = fun_of_list (i + 1) xs (g(i:=x))" + by (auto simp add: fun_eq_iff fun_of_list_def nth_Cons' + nat_diff_distrib [of "int (Suc 0)", simplified, symmetric] + diff_diff_eq) + +lemma nth_fun_of_list_eq: + "0 \ i \ i < int (length xs) \ xs ! nat i = fun_of_list 0 xs g i" + by (simp add: fun_of_list_def) + +end diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/RMD_Specification.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/RMD_Specification.thy Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,46 @@ +(* Title: HOL/SPARK/Examples/RIPEMD-160/RMD_Specification.thy + Author: Fabian Immler, TU Muenchen + +Verification of the RIPEMD-160 hash function +*) + +theory RMD_Specification +imports RMD SPARK +begin + +(* bit operations *) + +abbreviation rotate_left :: "int \ int \ int" where + "rotate_left i w == uint (word_rotl (nat i) (word_of_int w::word32))" + +spark_proof_functions + wordops__rotate_left = rotate_left + + +(* Conversions for proof functions *) +abbreviation k_l_spec :: " int => int " where + "k_l_spec j == uint (K (nat j))" +abbreviation k_r_spec :: " int => int " where + "k_r_spec j == uint (K' (nat j))" +abbreviation r_l_spec :: " int => int " where + "r_l_spec j == int (r (nat j))" +abbreviation r_r_spec :: " int => int " where + "r_r_spec j == int (r' (nat j))" +abbreviation s_l_spec :: " int => int " where + "s_l_spec j == int (s (nat j))" +abbreviation s_r_spec :: " int => int " where + "s_r_spec j == int (s' (nat j))" +abbreviation f_spec :: "int \ int \ int \ int \ int" where + "f_spec j x y z == + uint (f (nat j) (word_of_int x::word32) (word_of_int y) (word_of_int z))" + +spark_proof_functions + k_l_spec = k_l_spec + k_r_spec = k_r_spec + r_l_spec = r_l_spec + r_r_spec = r_r_spec + s_l_spec = s_l_spec + s_r_spec = s_r_spec + f_spec = f_spec + +end diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/R_L.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/R_L.thy Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,30 @@ +(* Title: HOL/SPARK/Examples/RIPEMD-160/R_L.thy + Author: Fabian Immler, TU Muenchen + +Verification of the RIPEMD-160 hash function +*) + +theory R_L +imports RMD_Specification RMD_Lemmas +begin + +spark_open "rmd/r_l.siv" + +spark_vc function_r_l_2 +proof - + from `0 \ j` `j \ 79` + show C: ?C1 + by (simp add: r_def r_list_def nth_map [symmetric, of _ _ int] del: fun_upd_apply) + (simp add: nth_fun_of_list_eq [of _ _ undefined] del: fun_upd_apply) + from C show ?C2 by simp + have "list_all (\n. int n \ 15) r_list" + by (simp add: r_list_def) + moreover have "length r_list = 80" + by (simp add: r_list_def) + ultimately show ?C3 unfolding C using `j \ 79` + by (simp add: r_def list_all_length) +qed + +spark_end + +end diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/R_R.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/R_R.thy Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,30 @@ +(* Title: HOL/SPARK/Examples/RIPEMD-160/R_R.thy + Author: Fabian Immler, TU Muenchen + +Verification of the RIPEMD-160 hash function +*) + +theory R_R +imports RMD_Specification RMD_Lemmas +begin + +spark_open "rmd/r_r.siv" + +spark_vc function_r_r_2 +proof - + from `0 \ j` `j \ 79` + show C: ?C1 + by (simp add: r'_def r'_list_def nth_map [symmetric, of _ _ int] del: fun_upd_apply) + (simp add: nth_fun_of_list_eq [of _ _ undefined] del: fun_upd_apply) + from C show ?C2 by simp + have "list_all (\n. int n \ 15) r'_list" + by (simp add: r'_list_def) + moreover have "length r'_list = 80" + by (simp add: r'_list_def) + ultimately show ?C3 unfolding C using `j \ 79` + by (simp add: r'_def list_all_length) +qed + +spark_end + +end diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/Round.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/Round.thy Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,465 @@ +(* Title: HOL/SPARK/Examples/RIPEMD-160/Round.thy + Author: Fabian Immler, TU Muenchen + +Verification of the RIPEMD-160 hash function +*) + +theory Round +imports RMD_Specification +begin + +spark_open "rmd/round.siv" + +abbreviation from_chain :: "chain \ RMD.chain" where + "from_chain c \ ( + word_of_int (h0 c), + word_of_int (h1 c), + word_of_int (h2 c), + word_of_int (h3 c), + word_of_int (h4 c))" + +abbreviation from_chain_pair :: "chain_pair \ RMD.chain \ RMD.chain" where + "from_chain_pair cc \ ( + from_chain (left cc), + from_chain (right cc))" + +abbreviation to_chain :: "RMD.chain \ chain" where + "to_chain c \ + (let (h0, h1, h2, h3, h4) = c in + (|h0 = uint h0, + h1 = uint h1, + h2 = uint h2, + h3 = uint h3, + h4 = uint h4|))" + +abbreviation to_chain_pair :: "RMD.chain \ RMD.chain \ chain_pair" where + "to_chain_pair c == (let (c1, c2) = c in + (| left = to_chain c1, + right = to_chain c2 |))" + +abbreviation steps' :: "chain_pair \ int \ block \ chain_pair" where + "steps' cc i b == to_chain_pair (steps + (\n. word_of_int (b (int n))) + (from_chain_pair cc) + (nat i))" + +abbreviation round_spec :: "chain \ block \ chain" where + "round_spec c b == to_chain (round (\n. word_of_int (b (int n))) (from_chain c))" + +spark_proof_functions + steps = steps' + round_spec = round_spec + +lemma uint_word_of_int_id: + assumes "0 <= (x::int)" + assumes "x <= 4294967295" + shows"uint(word_of_int x::word32) = x" + unfolding int_word_uint + using assms + by (simp add:int_mod_eq') + +lemma steps_step: "steps X cc (Suc i) = step_both X (steps X cc i) i" + unfolding steps_def + by (induct i) simp_all + +lemma from_to_id: "from_chain_pair (to_chain_pair CC) = CC" +proof (cases CC) + fix a::RMD.chain + fix b c d e f::word32 + assume "CC = (a, b, c, d, e, f)" + thus ?thesis by (cases a) simp +qed + +lemma steps_to_steps': + "F A (steps X cc i) B = + F A (from_chain_pair (to_chain_pair (steps X cc i))) B" + unfolding from_to_id .. + +lemma steps'_step: + assumes "0 <= i" + shows + "steps' cc (i + 1) X = to_chain_pair ( + step_both + (\n. word_of_int (X (int n))) + (from_chain_pair (steps' cc i X)) + (nat i))" +proof - + have "nat (i + 1) = Suc (nat i)" using assms by simp + show ?thesis + unfolding `nat (i + 1) = Suc (nat i)` steps_step steps_to_steps' + .. +qed + +lemma step_from_hyp: + assumes + step_hyp: + "\left = + \h0 = a, h1 = b, h2 = c, h3 = d, h4 = e\, + right = + \h0 = a', h1 = b', h2 = c', h3 = d', h4 = e'\\ = + steps' + (\left = + \h0 = a_0, h1 = b_0, h2 = c_0, + h3 = d_0, h4 = e_0\, + right = + \h0 = a_0, h1 = b_0, h2 = c_0, + h3 = d_0, h4 = e_0\\) + j x" + assumes "a <= 4294967295" (is "_ <= ?M") + assumes "b <= ?M" and "c <= ?M" and "d <= ?M" and "e <= ?M" + assumes "a' <= ?M" and "b' <= ?M" and "c' <= ?M" and "d' <= ?M" and "e' <= ?M" + assumes "0 <= a " and "0 <= b " and "0 <= c " and "0 <= d " and "0 <= e " + assumes "0 <= a'" and "0 <= b'" and "0 <= c'" and "0 <= d'" and "0 <= e'" + assumes "0 <= x (r_l_spec j)" and "x (r_l_spec j) <= ?M" + assumes "0 <= x (r_r_spec j)" and "x (r_r_spec j) <= ?M" + assumes "0 <= j" and "j <= 79" + shows + "\left = + \h0 = e, + h1 = + (rotate_left (s_l_spec j) + ((((a + f_spec j b c d) mod 4294967296 + + x (r_l_spec j)) mod + 4294967296 + + k_l_spec j) mod + 4294967296) + + e) mod + 4294967296, + h2 = b, h3 = rotate_left 10 c, + h4 = d\, + right = + \h0 = e', + h1 = + (rotate_left (s_r_spec j) + ((((a' + f_spec (79 - j) b' c' d') mod + 4294967296 + + x (r_r_spec j)) mod + 4294967296 + + k_r_spec j) mod + 4294967296) + + e') mod + 4294967296, + h2 = b', h3 = rotate_left 10 c', + h4 = d'\\ = + steps' + (\left = + \h0 = a_0, h1 = b_0, h2 = c_0, + h3 = d_0, h4 = e_0\, + right = + \h0 = a_0, h1 = b_0, h2 = c_0, + h3 = d_0, h4 = e_0\\) + (j + 1) x" + using step_hyp +proof - + let ?MM = 4294967296 + have AL: "uint(word_of_int e::word32) = e" + by (rule uint_word_of_int_id[OF `0 <= e` `e <= ?M`]) + have CL: "uint(word_of_int b::word32) = b" + by (rule uint_word_of_int_id[OF `0 <= b` `b <= ?M`]) + have DL: "True" .. + have EL: "uint(word_of_int d::word32) = d" + by (rule uint_word_of_int_id[OF `0 <= d` `d <= ?M`]) + have AR: "uint(word_of_int e'::word32) = e'" + by (rule uint_word_of_int_id[OF `0 <= e'` `e' <= ?M`]) + have CR: "uint(word_of_int b'::word32) = b'" + by (rule uint_word_of_int_id[OF `0 <= b'` `b' <= ?M`]) + have DR: "True" .. + have ER: "uint(word_of_int d'::word32) = d'" + by (rule uint_word_of_int_id[OF `0 <= d'` `d' <= ?M`]) + have BL: + "(uint + (word_rotl (s (nat j)) + ((word_of_int::int\word32) + ((((a + f_spec j b c d) mod ?MM + + x (r_l_spec j)) mod ?MM + + k_l_spec j) mod ?MM))) + + e) mod ?MM + = + uint + (word_rotl (s (nat j)) + (word_of_int a + + f (nat j) (word_of_int b) + (word_of_int c) (word_of_int d) + + word_of_int (x (r_l_spec j)) + + K (nat j)) + + word_of_int e)" + (is "(uint (word_rotl _ (_ ((((_ + ?F) mod _ + ?X) mod _ + _) mod _))) + _) mod _ = _") + proof - + have "a mod ?MM = a" using `0 <= a` `a <= ?M` + by (simp add: int_mod_eq') + have "?X mod ?MM = ?X" using `0 <= ?X` `?X <= ?M` + by (simp add: int_mod_eq') + have "e mod ?MM = e" using `0 <= e` `e <= ?M` + by (simp add: int_mod_eq') + have "(?MM::int) = 2 ^ len_of TYPE(32)" by simp + show ?thesis + unfolding + word_add_alt + uint_word_of_int_id[OF `0 <= a` `a <= ?M`] + uint_word_of_int_id[OF `0 <= ?X` `?X <= ?M`] + int_word_uint + unfolding `?MM = 2 ^ len_of TYPE(32)` + unfolding word_uint.Abs_norm + by (simp add: + `a mod ?MM = a` + `e mod ?MM = e` + `?X mod ?MM = ?X`) + qed + + have BR: + "(uint + (word_rotl (s' (nat j)) + ((word_of_int::int\word32) + ((((a' + f_spec (79 - j) b' c' d') mod ?MM + + x (r_r_spec j)) mod ?MM + + k_r_spec j) mod ?MM))) + + e') mod ?MM + = + uint + (word_rotl (s' (nat j)) + (word_of_int a' + + f (79 - nat j) (word_of_int b') + (word_of_int c') (word_of_int d') + + word_of_int (x (r_r_spec j)) + + K' (nat j)) + + word_of_int e')" + (is "(uint (word_rotl _ (_ ((((_ + ?F) mod _ + ?X) mod _ + _) mod _))) + _) mod _ = _") + proof - + have "a' mod ?MM = a'" using `0 <= a'` `a' <= ?M` + by (simp add: int_mod_eq') + have "?X mod ?MM = ?X" using `0 <= ?X` `?X <= ?M` + by (simp add: int_mod_eq') + have "e' mod ?MM = e'" using `0 <= e'` `e' <= ?M` + by (simp add: int_mod_eq') + have "(?MM::int) = 2 ^ len_of TYPE(32)" by simp + have nat_transfer: "79 - nat j = nat (79 - j)" + using nat_diff_distrib `0 <= j` `j <= 79` + by simp + show ?thesis + unfolding + word_add_alt + uint_word_of_int_id[OF `0 <= a'` `a' <= ?M`] + uint_word_of_int_id[OF `0 <= ?X` `?X <= ?M`] + int_word_uint + nat_transfer + unfolding `?MM = 2 ^ len_of TYPE(32)` + unfolding word_uint.Abs_norm + by (simp add: + `a' mod ?MM = a'` + `e' mod ?MM = e'` + `?X mod ?MM = ?X`) + qed + + show ?thesis + unfolding steps'_step[OF `0 <= j`] step_hyp[symmetric] + step_both_def step_r_def step_l_def + by (simp add: AL BL CL DL EL AR BR CR DR ER) +qed + +spark_vc procedure_round_61 +proof - + let ?M = "4294967295::int" + have step_hyp: + "\left = + \h0 = ca, h1 = cb, h2 = cc, + h3 = cd, h4 = ce\, + right = + \h0 = ca, h1 = cb, h2 = cc, + h3 = cd, h4 = ce\\ = + steps' + (\left = + \h0 = ca, h1 = cb, h2 = cc, + h3 = cd, h4 = ce\, + right = + \h0 = ca, h1 = cb, h2 = cc, + h3 = cd, h4 = ce\\) + 0 x" + unfolding steps_def + by (simp add: + uint_word_of_int_id[OF `0 <= ca` `ca <= ?M`] + uint_word_of_int_id[OF `0 <= cb` `cb <= ?M`] + uint_word_of_int_id[OF `0 <= cc` `cc <= ?M`] + uint_word_of_int_id[OF `0 <= cd` `cd <= ?M`] + uint_word_of_int_id[OF `0 <= ce` `ce <= ?M`]) + let ?rotate_arg_l = + "((((ca + f 0 cb cc cd) smod 4294967296 + + x (r_l 0)) smod 4294967296 + k_l 0) smod 4294967296)" + let ?rotate_arg_r = + "((((ca + f 79 cb cc cd) smod 4294967296 + + x (r_r 0)) smod 4294967296 + k_r 0) smod 4294967296)" + note returns = + `wordops__rotate (s_l 0) ?rotate_arg_l = + rotate_left (s_l 0) ?rotate_arg_l` + `wordops__rotate (s_r 0) ?rotate_arg_r = + rotate_left (s_r 0) ?rotate_arg_r` + `wordops__rotate 10 cc = rotate_left 10 cc` + `f 0 cb cc cd = f_spec 0 cb cc cd` + `f 79 cb cc cd = f_spec 79 cb cc cd` + `k_l 0 = k_l_spec 0` + `k_r 0 = k_r_spec 0` + `r_l 0 = r_l_spec 0` + `r_r 0 = r_r_spec 0` + `s_l 0 = s_l_spec 0` + `s_r 0 = s_r_spec 0` + + note x_borders = `\i. 0 \ i \ i \ 15 \ 0 \ x i \ x i \ ?M` + + from `0 <= r_l 0` `r_l 0 <= 15` x_borders + have "0 \ x (r_l 0)" by blast + hence x_lower: "0 <= x (r_l_spec 0)" unfolding returns . + + from `0 <= r_l 0` `r_l 0 <= 15` x_borders + have "x (r_l 0) <= ?M" by blast + hence x_upper: "x (r_l_spec 0) <= ?M" unfolding returns . + + from `0 <= r_r 0` `r_r 0 <= 15` x_borders + have "0 \ x (r_r 0)" by blast + hence x_lower': "0 <= x (r_r_spec 0)" unfolding returns . + + from `0 <= r_r 0` `r_r 0 <= 15` x_borders + have "x (r_r 0) <= ?M" by blast + hence x_upper': "x (r_r_spec 0) <= ?M" unfolding returns . + + have "0 <= (0::int)" by simp + have "0 <= (79::int)" by simp + note step_from_hyp [OF + step_hyp + H2 H4 H6 H8 H10 H2 H4 H6 H8 H10 (* upper bounds *) + H1 H3 H5 H7 H9 H1 H3 H5 H7 H9 (* lower bounds *) + ] + from this[OF x_lower x_upper x_lower' x_upper' `0 <= 0` `0 <= 79`] + `0 \ ca` `0 \ ce` x_lower x_lower' + show ?thesis unfolding returns(1) returns(2) unfolding returns + by (simp add: smod_pos_pos) +qed + +spark_vc procedure_round_62 +proof - + let ?M = "4294967295::int" + let ?rotate_arg_l = + "((((cla + f (loop__1__j + 1) clb clc cld) smod 4294967296 + + x (r_l (loop__1__j + 1))) smod 4294967296 + + k_l (loop__1__j + 1)) smod 4294967296)" + let ?rotate_arg_r = + "((((cra + f (79 - (loop__1__j + 1)) crb crc crd) smod + 4294967296 + x (r_r (loop__1__j + 1))) smod 4294967296 + + k_r (loop__1__j + 1)) smod 4294967296)" + + have s: "78 - loop__1__j = (79 - (loop__1__j + 1))" by simp + note returns = + `wordops__rotate (s_l (loop__1__j + 1)) ?rotate_arg_l = + rotate_left (s_l (loop__1__j + 1)) ?rotate_arg_l` + `wordops__rotate (s_r (loop__1__j + 1)) ?rotate_arg_r = + rotate_left (s_r (loop__1__j + 1)) ?rotate_arg_r` + `f (loop__1__j + 1) clb clc cld = + f_spec (loop__1__j + 1) clb clc cld` + `f (78 - loop__1__j) crb crc crd = + f_spec (78 - loop__1__j) crb crc crd`[simplified s] + `wordops__rotate 10 clc = rotate_left 10 clc` + `wordops__rotate 10 crc = rotate_left 10 crc` + `k_l (loop__1__j + 1) = k_l_spec (loop__1__j + 1)` + `k_r (loop__1__j + 1) = k_r_spec (loop__1__j + 1)` + `r_l (loop__1__j + 1) = r_l_spec (loop__1__j + 1)` + `r_r (loop__1__j + 1) = r_r_spec (loop__1__j + 1)` + `s_l (loop__1__j + 1) = s_l_spec (loop__1__j + 1)` + `s_r (loop__1__j + 1) = s_r_spec (loop__1__j + 1)` + + note x_borders = `\i. 0 \ i \ i \ 15 \ 0 \ x i \ x i \ ?M` + + from `0 <= r_l (loop__1__j + 1)` `r_l (loop__1__j + 1) <= 15` x_borders + have "0 \ x (r_l (loop__1__j + 1))" by blast + hence x_lower: "0 <= x (r_l_spec (loop__1__j + 1))" unfolding returns . + + from `0 <= r_l (loop__1__j + 1)` `r_l (loop__1__j + 1) <= 15` x_borders + have "x (r_l (loop__1__j + 1)) <= ?M" by blast + hence x_upper: "x (r_l_spec (loop__1__j + 1)) <= ?M" unfolding returns . + + from `0 <= r_r (loop__1__j + 1)` `r_r (loop__1__j + 1) <= 15` x_borders + have "0 \ x (r_r (loop__1__j + 1))" by blast + hence x_lower': "0 <= x (r_r_spec (loop__1__j + 1))" unfolding returns . + + from `0 <= r_r (loop__1__j + 1)` `r_r (loop__1__j + 1) <= 15` x_borders + have "x (r_r (loop__1__j + 1)) <= ?M" by blast + hence x_upper': "x (r_r_spec (loop__1__j + 1)) <= ?M" unfolding returns . + + from `0 <= loop__1__j` have "0 <= loop__1__j + 1" by simp + from `loop__1__j <= 78` have "loop__1__j + 1 <= 79" by simp + + have "loop__1__j + 1 + 1 = loop__1__j + 2" by simp + + note step_from_hyp[OF H1 + `cla <= ?M` + `clb <= ?M` + `clc <= ?M` + `cld <= ?M` + `cle <= ?M` + `cra <= ?M` + `crb <= ?M` + `crc <= ?M` + `crd <= ?M` + `cre <= ?M` + + `0 <= cla` + `0 <= clb` + `0 <= clc` + `0 <= cld` + `0 <= cle` + `0 <= cra` + `0 <= crb` + `0 <= crc` + `0 <= crd` + `0 <= cre`] + from this[OF + x_lower x_upper x_lower' x_upper' + `0 <= loop__1__j + 1` `loop__1__j + 1 <= 79`] + `0 \ cla` `0 \ cle` `0 \ cra` `0 \ cre` x_lower x_lower' + show ?thesis unfolding `loop__1__j + 1 + 1 = loop__1__j + 2` + unfolding returns(1) returns(2) unfolding returns + by (simp add: smod_pos_pos) +qed + +spark_vc procedure_round_76 +proof - + let ?M = "4294967295 :: int" + let ?INIT_CHAIN = + "\h0 = ca_init, h1 = cb_init, + h2 = cc_init, h3 = cd_init, + h4 = ce_init\" + have steps_to_steps': + "steps + (\n\nat. word_of_int (x (int n))) + (from_chain ?INIT_CHAIN, from_chain ?INIT_CHAIN) + 80 = + from_chain_pair ( + steps' + (\left = ?INIT_CHAIN, right = ?INIT_CHAIN\) + 80 + x)" + unfolding from_to_id by simp + from + `0 \ ca_init` `ca_init \ ?M` + `0 \ cb_init` `cb_init \ ?M` + `0 \ cc_init` `cc_init \ ?M` + `0 \ cd_init` `cd_init \ ?M` + `0 \ ce_init` `ce_init \ ?M` + `0 \ cla` `cla \ ?M` + `0 \ clb` `clb \ ?M` + `0 \ clc` `clc \ ?M` + `0 \ cld` `cld \ ?M` + `0 \ cle` `cle \ ?M` + `0 \ cra` `cra \ ?M` + `0 \ crb` `crb \ ?M` + `0 \ crc` `crc \ ?M` + `0 \ crd` `crd \ ?M` + `0 \ cre` `cre \ ?M` + show ?thesis + unfolding round_def + unfolding steps_to_steps' + unfolding H1[symmetric] + by (simp add: uint_word_ariths(2) rdmods smod_pos_pos + uint_word_of_int_id) +qed + +spark_end + +end diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/S_L.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/S_L.thy Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,30 @@ +(* Title: HOL/SPARK/Examples/RIPEMD-160/S_L.thy + Author: Fabian Immler, TU Muenchen + +Verification of the RIPEMD-160 hash function +*) + +theory S_L +imports RMD_Specification RMD_Lemmas +begin + +spark_open "rmd/s_l.siv" + +spark_vc function_s_l_2 +proof - + from `0 \ j` `j \ 79` + show C: ?C1 + by (simp add: s_def s_list_def nth_map [symmetric, of _ _ int] del: fun_upd_apply) + (simp add: nth_fun_of_list_eq [of _ _ undefined] del: fun_upd_apply) + from C show ?C2 by simp + have "list_all (\n. int n \ 15) s_list" + by (simp add: s_list_def) + moreover have "length s_list = 80" + by (simp add: s_list_def) + ultimately show ?C3 unfolding C using `j \ 79` + by (simp add: s_def list_all_length) +qed + +spark_end + +end \ No newline at end of file diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/S_R.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/S_R.thy Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,30 @@ +(* Title: HOL/SPARK/Examples/RIPEMD-160/S_R.thy + Author: Fabian Immler, TU Muenchen + +Verification of the RIPEMD-160 hash function +*) + +theory S_R +imports RMD_Specification RMD_Lemmas +begin + +spark_open "rmd/s_r.siv" + +spark_vc function_s_r_2 +proof - + from `0 \ j` `j \ 79` + show C: ?C1 + by (simp add: s'_def s'_list_def nth_map [symmetric, of _ _ int] del: fun_upd_apply) + (simp add: nth_fun_of_list_eq [of _ _ undefined] del: fun_upd_apply) + from C show ?C2 by simp + have "list_all (\n. int n \ 15) s'_list" + by (simp add: s'_list_def) + moreover have "length s'_list = 80" + by (simp add: s'_list_def) + ultimately show ?C3 unfolding C using `j \ 79` + by (simp add: s'_def list_all_length) +qed + +spark_end + +end \ No newline at end of file diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd.adb --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd.adb Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,187 @@ +package body RMD is + + + + function F(J : Round_Index; X,Y,Z : Word) return Word + is + Result: Word; + begin + if 0 <= J and J <= 15 then Result := X xor Y xor Z; + elsif 16 <= J and J <= 31 then Result := (X and Y) or (not X and Z); + elsif 32 <= J and J <= 47 then Result := (X or not Y) xor Z; + elsif 48 <= J and J <= 63 then Result := (X and Z) or (Y and not Z); + else Result := X xor (Y or not Z); + end if; + return Result; + end F; + + + + function K_L(J : Round_Index) return Word + is + K: Word; + begin + if 0 <= J and J <= 15 then K := 16#0000_0000#; + elsif 16 <= J and J <= 31 then K := 16#5A82_7999#; + elsif 32 <= J and J <= 47 then K := 16#6ED9_EBA1#; + elsif 48 <= J and J <= 63 then K := 16#8F1B_BCDC#; + else K := 16#A953_FD4E#; + end if; + return K; + end K_L; + + + function K_R(J : Round_Index) return Word + is + K: Word; + begin + if 0 <= J and J <= 15 then K := 16#50A2_8BE6#; + elsif 16 <= J and J <= 31 then K := 16#5C4D_D124#; + elsif 32 <= J and J <= 47 then K := 16#6D70_3EF3#; + elsif 48 <= J and J <= 63 then K := 16#7A6D_76E9#; + else K := 16#0000_0000#; + end if; + return K; + end K_R; + + + + function R_L(J : Round_Index) return Block_Index + is + R_Values : constant Block_Permutation := Block_Permutation' + (0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, + 7, 4, 13, 1, 10, 6, 15, 3, 12, 0, 9, 5, 2, 14, 11, 8, + 3, 10, 14, 4, 9, 15, 8, 1, 2, 7, 0, 6, 13, 11, 5, 12, + 1, 9, 11, 10, 0, 8, 12, 4, 13, 3, 7, 15, 14, 5, 6, 2, + 4, 0, 5, 9, 7, 12, 2, 10, 14, 1, 3, 8, 11, 6, 15, 13); + --# for R_Values declare rule; + begin + return R_Values(J); + end R_L; + + + function R_R(J : Round_Index) return Block_Index + is + R_Values : constant Block_Permutation := Block_Permutation' + (5, 14, 7, 0, 9, 2, 11, 4, 13, 6, 15, 8, 1, 10, 3, 12, + 6, 11, 3, 7, 0, 13, 5, 10, 14, 15, 8, 12, 4, 9, 1, 2, + 15, 5, 1, 3, 7, 14, 6, 9, 11, 8, 12, 2, 10, 0, 4, 13, + 8, 6, 4, 1, 3, 11, 15, 0, 5, 12, 2, 13, 9, 7, 10, 14, + 12, 15, 10, 4, 1, 5, 8, 7, 6, 2, 13, 14, 0, 3, 9, 11); + --# for R_Values declare rule; + begin + return R_Values(J); + end R_R; + + + function S_L(J : Round_Index) return Rotate_Amount + is + S_Values : constant Rotate_Definition := Rotate_Definition' + (11, 14, 15, 12, 5, 8, 7, 9, 11, 13, 14, 15, 6, 7, 9, 8, + 7, 6, 8, 13, 11, 9, 7, 15, 7, 12, 15, 9, 11, 7, 13, 12, + 11, 13, 6, 7, 14, 9, 13, 15, 14, 8, 13, 6, 5, 12, 7, 5, + 11, 12, 14, 15, 14, 15, 9, 8, 9, 14, 5, 6, 8, 6, 5, 12, + 9, 15, 5, 11, 6, 8, 13, 12, 5, 12, 13, 14, 11, 8, 5, 6); + --# for S_Values declare rule; + begin + return S_Values(J); + end S_L; + + + function S_R(J : Round_Index) return Rotate_Amount + is + S_Values : constant Rotate_Definition := Rotate_Definition' + (8, 9, 9, 11, 13, 15, 15, 5, 7, 7, 8, 11, 14, 14, 12, 6, + 9, 13, 15, 7, 12, 8, 9, 11, 7, 7, 12, 7, 6, 15, 13, 11, + 9, 7, 15, 11, 8, 6, 6, 14, 12, 13, 5, 14, 13, 13, 7, 5, + 15, 5, 8, 11, 14, 14, 6, 14, 6, 9, 12, 9, 12, 5, 15, 8, + 8, 5, 12, 9, 12, 5, 14, 6, 8, 13, 6, 5, 15, 13, 11, 11); + --# for S_Values declare rule; + begin + return S_Values(J); + end S_R; + + + + procedure Round(CA, CB, CC, CD, CE : in out Word; X : in Block) + is + CLA, CLB, CLC, CLD, CLE, CRA, CRB, CRC, CRD, CRE : Word; + T : Word; + begin + CLA := CA; + CLB := CB; + CLC := CC; + CLD := CD; + CLE := CE; + CRA := CA; + CRB := CB; + CRC := CC; + CRD := CD; + CRE := CE; + for J in Round_Index range 0..79 + loop + -- left + T := Wordops.Rotate(S_L(J), + CLA + + F(J, CLB, CLC, CLD) + + X(R_L(J)) + + K_L(J)) + + CLE; + CLA := CLE; + CLE := CLD; + CLD := Wordops.Rotate(10, CLC); + CLC := CLB; + CLB := T; + -- right + T := Wordops.Rotate(S_R(J), + CRA + + F(79 - J, CRB, CRC, CRD) + + X(R_R(J)) + + K_R(J)) + + CRE; + CRA := CRE; + CRE := CRD; + CRD := Wordops.Rotate(10, CRC); + CRC := CRB; + CRB := T; + --# assert Chain_Pair'(Chain'(CLA, CLB, CLC, CLD, CLE), + --# Chain'(CRA, CRB, CRC, CRD, CRE)) = + --# steps(Chain_Pair'(Chain'(CA~, CB~, CC~, CD~, CE~), + --# Chain'(CA~, CB~, CC~, CD~, CE~)), J + 1, X) + --# and CA = CA~ and CB = CB~ and CC = CC~ and CD = CD~ and CE = CE~; + end loop; + T := CB + CLC + CRD; + CB := CC + CLD + CRE; + CC := CD + CLE + CRA; + CD := CE + CLA + CRB; + CE := CA + CLB + CRC; + CA := T; + end Round; + + function Hash(X : Message) return Chain + is + CA_Init : constant Word := 16#6745_2301#; + CB_Init : constant Word := 16#EFCD_AB89#; + CC_Init : constant Word := 16#98BA_DCFE#; + CD_Init : constant Word := 16#1032_5476#; + CE_Init : constant Word := 16#C3D2_E1F0#; + CA, CB, CC, CD, CE : Word; + begin + CA := CA_Init; + CB := CB_Init; + CC := CC_Init; + CD := CD_Init; + CE := CE_Init; + for I in Message_Index range X'First..X'Last + loop + Round(CA, CB, CC, CD, CE, X(I)); + --# assert Chain'(CA, CB, CC, CD, CE) = rounds( + --# Chain'(CA_Init, CB_Init, CC_Init, CD_Init, CE_Init), + --# I + 1, + --# X); + end loop; + return Chain'(CA, CB, CC, CD, CE); + end Hash; + +end RMD; + diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd.ads --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd.ads Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,92 @@ +with Wordops; +use type Wordops.Word; +--# inherit Wordops; + +package RMD +is + + -- Types + + subtype Word is Wordops.Word; + + type Chain is + record + H0, H1, H2, H3, H4 : Word; + end record; + + type Block_Index is range 0..15; + type Block is array(Block_Index) of Word; + + type Message_Index is range 0..2**32; + type Message is array(Message_Index range <>) of Block; + + -- Isabelle specification + + --# function rmd_hash(X : Message; L : Message_Index) return Chain; + + function Hash(X : Message) return Chain; + --# pre X'First = 0; + --# return rmd_hash(X, X'Last + 1); + +private + + -- Types + + type Round_Index is range 0..79; + + type Chain_Pair is + record + Left, Right : Chain; + end record; + + type Block_Permutation is array(Round_Index) of Block_Index; + + subtype Rotate_Amount is Wordops.Rotate_Amount; + type Rotate_Definition is array(Round_Index) of Rotate_Amount; + + + -- Isabelle proof functions + + --# function f_spec(J : Round_Index; X,Y,Z : Word) return Word; + --# function K_l_spec(J : Round_Index) return Word; + --# function K_r_spec(J : Round_Index) return Word; + --# function r_l_spec(J : Round_Index) return Block_Index; + --# function r_r_spec(J : Round_Index) return Block_Index; + --# function s_l_spec(J : Round_Index) return Rotate_Amount; + --# function s_r_spec(J : Round_Index) return Rotate_Amount; + --# function steps(CS : Chain_Pair; I : Round_Index; B : Block) + --# return Chain_Pair; + --# function round_spec(C : Chain; B : Block) return Chain; + --# function rounds(C : Chain; I : Message_Index; X : Message) + --# return Chain; + + + -- Spark Implementation + + function F(J : Round_Index; X,Y,Z : Word) return Word; + --# return f_spec(J, X, Y, Z); + + function K_L(J : Round_Index) return Word; + --# return K_l_spec(J); + + function K_R(J : Round_Index) return Word; + --# return K_r_spec(J); + + function R_L(J : Round_Index) return Block_Index; + --# return r_l_spec(J); + + function R_R(J : Round_Index) return Block_Index; + --# return r_r_spec(J); + + function S_L(J : Round_Index) return Rotate_Amount; + --# return s_l_spec(J); + + function S_R(J : Round_Index) return Rotate_Amount; + --# return s_r_spec(J); + + procedure Round(CA, CB, CC, CD, CE : in out Word; X: in Block); + --# derives CA, CB, CC, CD, CE from X, CA, CB, CC, CD, CE; + --# post Chain'(CA, CB, CC, CD, CE) = + --# round_spec(Chain'(CA~, CB~, CC~, CD~, CE~), X); + +end RMD; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/f.fdl --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/f.fdl Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,41 @@ + {*******************************************************} + {FDL Declarations} + {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039} + {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.} + {*******************************************************} + + + {DATE : 29-NOV-2010 14:30:19.73} + + {function RMD.F} + + +title function f; + + function round__(real) : integer; + type interfaces__unsigned_32 = integer; + type round_index = integer; + const round_index__base__first : integer = pending; + const round_index__base__last : integer = pending; + const word__base__first : integer = pending; + const word__base__last : integer = pending; + const interfaces__unsigned_32__base__first : integer = pending; + const interfaces__unsigned_32__base__last : integer = pending; + const round_index__first : integer = pending; + const round_index__last : integer = pending; + const round_index__size : integer = pending; + const word__first : integer = pending; + const word__last : integer = pending; + const word__modulus : integer = pending; + const word__size : integer = pending; + const interfaces__unsigned_32__first : integer = pending; + const interfaces__unsigned_32__last : integer = pending; + const interfaces__unsigned_32__modulus : integer = pending; + const interfaces__unsigned_32__size : integer = pending; + var j : integer; + var x : integer; + var y : integer; + var z : integer; + function f_spec(integer, integer, integer, integer) : integer; + +end; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/f.rls --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/f.rls Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,35 @@ + /*********************************************************/ + /*Proof Rule Declarations*/ + /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/ + /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/ + /*********************************************************/ + + + /*DATE : 29-NOV-2010 14:30:19.73*/ + + /*function RMD.F*/ + + +rule_family f_rules: + X requires [X:any] & + X <= Y requires [X:ire, Y:ire] & + X >= Y requires [X:ire, Y:ire]. + +f_rules(1): interfaces__unsigned_32__size >= 0 may_be_deduced. +f_rules(2): interfaces__unsigned_32__first may_be_replaced_by 0. +f_rules(3): interfaces__unsigned_32__last may_be_replaced_by 4294967295. +f_rules(4): interfaces__unsigned_32__base__first may_be_replaced_by 0. +f_rules(5): interfaces__unsigned_32__base__last may_be_replaced_by 4294967295. +f_rules(6): interfaces__unsigned_32__modulus may_be_replaced_by 4294967296. +f_rules(7): word__size >= 0 may_be_deduced. +f_rules(8): word__first may_be_replaced_by 0. +f_rules(9): word__last may_be_replaced_by 4294967295. +f_rules(10): word__base__first may_be_replaced_by 0. +f_rules(11): word__base__last may_be_replaced_by 4294967295. +f_rules(12): word__modulus may_be_replaced_by 4294967296. +f_rules(13): round_index__size >= 0 may_be_deduced. +f_rules(14): round_index__first may_be_replaced_by 0. +f_rules(15): round_index__last may_be_replaced_by 79. +f_rules(16): round_index__base__first <= round_index__base__last may_be_deduced. +f_rules(17): round_index__base__first <= round_index__first may_be_deduced. +f_rules(18): round_index__base__last >= round_index__last may_be_deduced. diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/f.siv --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/f.siv Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,228 @@ +***************************************************************************** + Semantic Analysis of SPARK Text + Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 + Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. +***************************************************************************** + + +CREATED 29-NOV-2010, 14:30:19 SIMPLIFIED 29-NOV-2010, 14:30:28 + +SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 +Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. + +function RMD.F + + + + +For path(s) from start to run-time check associated with statement of line 9: + +function_f_1. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 10: + +function_f_2. +H1: x >= 0 . +H2: x <= 4294967295 . +H3: y >= 0 . +H4: y <= 4294967295 . +H5: z >= 0 . +H6: z <= 4294967295 . +H7: 16 <= j . +H8: j <= 31 . +H9: interfaces__unsigned_32__size >= 0 . +H10: word__size >= 0 . +H11: round_index__size >= 0 . +H12: round_index__base__first <= round_index__base__last . +H13: round_index__base__first <= 0 . +H14: round_index__base__last >= 79 . + -> +C1: bit__or(bit__and(x, y), bit__and(4294967295 - x, z)) >= 0 . +C2: bit__or(bit__and(x, y), bit__and(4294967295 - x, z)) <= 4294967295 . + + +For path(s) from start to run-time check associated with statement of line 11: + +function_f_3. +H1: x >= 0 . +H2: x <= 4294967295 . +H3: y >= 0 . +H4: y <= 4294967295 . +H5: z >= 0 . +H6: z <= 4294967295 . +H7: 32 <= j . +H8: j <= 47 . +H9: interfaces__unsigned_32__size >= 0 . +H10: word__size >= 0 . +H11: round_index__size >= 0 . +H12: round_index__base__first <= round_index__base__last . +H13: round_index__base__first <= 0 . +H14: round_index__base__last >= 79 . + -> +C1: bit__xor(bit__or(x, 4294967295 - y), z) >= 0 . +C2: bit__xor(bit__or(x, 4294967295 - y), z) <= 4294967295 . + + +For path(s) from start to run-time check associated with statement of line 12: + +function_f_4. +H1: x >= 0 . +H2: x <= 4294967295 . +H3: y >= 0 . +H4: y <= 4294967295 . +H5: z >= 0 . +H6: z <= 4294967295 . +H7: 48 <= j . +H8: j <= 63 . +H9: interfaces__unsigned_32__size >= 0 . +H10: word__size >= 0 . +H11: round_index__size >= 0 . +H12: round_index__base__first <= round_index__base__last . +H13: round_index__base__first <= 0 . +H14: round_index__base__last >= 79 . + -> +C1: bit__or(bit__and(x, z), bit__and(y, 4294967295 - z)) >= 0 . +C2: bit__or(bit__and(x, z), bit__and(y, 4294967295 - z)) <= 4294967295 . + + +For path(s) from start to run-time check associated with statement of line 13: + +function_f_5. +H1: j >= 0 . +H2: j <= 79 . +H3: x >= 0 . +H4: x <= 4294967295 . +H5: y >= 0 . +H6: y <= 4294967295 . +H7: z >= 0 . +H8: z <= 4294967295 . +H9: 15 < j . +H10: 31 < j . +H11: 47 < j . +H12: 63 < j . +H13: interfaces__unsigned_32__size >= 0 . +H14: word__size >= 0 . +H15: round_index__size >= 0 . +H16: round_index__base__first <= round_index__base__last . +H17: round_index__base__first <= 0 . +H18: round_index__base__last >= 79 . + -> +C1: bit__xor(x, bit__or(y, 4294967295 - z)) >= 0 . +C2: bit__xor(x, bit__or(y, 4294967295 - z)) <= 4294967295 . + + +For path(s) from start to finish: + +function_f_6. +H1: j >= 0 . +H2: x >= 0 . +H3: x <= 4294967295 . +H4: y >= 0 . +H5: y <= 4294967295 . +H6: z >= 0 . +H7: z <= 4294967295 . +H8: j <= 15 . +H9: bit__xor(x, bit__xor(y, z)) >= 0 . +H10: bit__xor(x, bit__xor(y, z)) <= 4294967295 . +H11: interfaces__unsigned_32__size >= 0 . +H12: word__size >= 0 . +H13: round_index__size >= 0 . +H14: round_index__base__first <= round_index__base__last . +H15: round_index__base__first <= 0 . +H16: round_index__base__last >= 79 . + -> +C1: bit__xor(x, bit__xor(y, z)) = f_spec(j, x, y, z) . + + +function_f_7. +H1: x >= 0 . +H2: x <= 4294967295 . +H3: y >= 0 . +H4: y <= 4294967295 . +H5: z >= 0 . +H6: z <= 4294967295 . +H7: 16 <= j . +H8: j <= 31 . +H9: bit__or(bit__and(x, y), bit__and(4294967295 - x, z)) >= 0 . +H10: bit__or(bit__and(x, y), bit__and(4294967295 - x, z)) <= 4294967295 . +H11: interfaces__unsigned_32__size >= 0 . +H12: word__size >= 0 . +H13: round_index__size >= 0 . +H14: round_index__base__first <= round_index__base__last . +H15: round_index__base__first <= 0 . +H16: round_index__base__last >= 79 . + -> +C1: bit__or(bit__and(x, y), bit__and(4294967295 - x, z)) = f_spec(j, x, y, z) + . + + +function_f_8. +H1: x >= 0 . +H2: x <= 4294967295 . +H3: y >= 0 . +H4: y <= 4294967295 . +H5: z >= 0 . +H6: z <= 4294967295 . +H7: 32 <= j . +H8: j <= 47 . +H9: bit__xor(bit__or(x, 4294967295 - y), z) >= 0 . +H10: bit__xor(bit__or(x, 4294967295 - y), z) <= 4294967295 . +H11: interfaces__unsigned_32__size >= 0 . +H12: word__size >= 0 . +H13: round_index__size >= 0 . +H14: round_index__base__first <= round_index__base__last . +H15: round_index__base__first <= 0 . +H16: round_index__base__last >= 79 . + -> +C1: bit__xor(bit__or(x, 4294967295 - y), z) = f_spec(j, x, y, z) . + + +function_f_9. +H1: x >= 0 . +H2: x <= 4294967295 . +H3: y >= 0 . +H4: y <= 4294967295 . +H5: z >= 0 . +H6: z <= 4294967295 . +H7: 48 <= j . +H8: j <= 63 . +H9: bit__or(bit__and(x, z), bit__and(y, 4294967295 - z)) >= 0 . +H10: bit__or(bit__and(x, z), bit__and(y, 4294967295 - z)) <= 4294967295 . +H11: interfaces__unsigned_32__size >= 0 . +H12: word__size >= 0 . +H13: round_index__size >= 0 . +H14: round_index__base__first <= round_index__base__last . +H15: round_index__base__first <= 0 . +H16: round_index__base__last >= 79 . + -> +C1: bit__or(bit__and(x, z), bit__and(y, 4294967295 - z)) = f_spec(j, x, y, z) + . + + +function_f_10. +H1: j >= 0 . +H2: j <= 79 . +H3: x >= 0 . +H4: x <= 4294967295 . +H5: y >= 0 . +H6: y <= 4294967295 . +H7: z >= 0 . +H8: z <= 4294967295 . +H9: 15 < j . +H10: 31 < j . +H11: 47 < j . +H12: 63 < j . +H13: bit__xor(x, bit__or(y, 4294967295 - z)) >= 0 . +H14: bit__xor(x, bit__or(y, 4294967295 - z)) <= 4294967295 . +H15: interfaces__unsigned_32__size >= 0 . +H16: word__size >= 0 . +H17: round_index__size >= 0 . +H18: round_index__base__first <= round_index__base__last . +H19: round_index__base__first <= 0 . +H20: round_index__base__last >= 79 . + -> +C1: bit__xor(x, bit__or(y, 4294967295 - z)) = f_spec(j, x, y, z) . + + diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/hash.fdl --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/hash.fdl Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,74 @@ + {*******************************************************} + {FDL Declarations} + {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039} + {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.} + {*******************************************************} + + + {DATE : 29-NOV-2010 14:30:20.17} + + {function RMD.Hash} + + +title function hash; + + function round__(real) : integer; + type interfaces__unsigned_32 = integer; + type block_index = integer; + type message_index = integer; + type chain = record + h0 : integer; + h1 : integer; + h2 : integer; + h3 : integer; + h4 : integer + end; + type block = array [integer] of integer; + type message = array [integer] of block; + const ca_init : integer = pending; + const cb_init : integer = pending; + const cc_init : integer = pending; + const cd_init : integer = pending; + const ce_init : integer = pending; + const message_index__base__first : integer = pending; + const message_index__base__last : integer = pending; + const block_index__base__first : integer = pending; + const block_index__base__last : integer = pending; + const word__base__first : integer = pending; + const word__base__last : integer = pending; + const interfaces__unsigned_32__base__first : integer = pending; + const interfaces__unsigned_32__base__last : integer = pending; + const x__index__subtype__1__first : integer = pending; + const x__index__subtype__1__last : integer = pending; + const message_index__first : integer = pending; + const message_index__last : integer = pending; + const message_index__size : integer = pending; + const block_index__first : integer = pending; + const block_index__last : integer = pending; + const block_index__size : integer = pending; + const chain__size : integer = pending; + const word__first : integer = pending; + const word__last : integer = pending; + const word__modulus : integer = pending; + const word__size : integer = pending; + const interfaces__unsigned_32__first : integer = pending; + const interfaces__unsigned_32__last : integer = pending; + const interfaces__unsigned_32__modulus : integer = pending; + const interfaces__unsigned_32__size : integer = pending; + var x : message; + var ca : integer; + var cb : integer; + var cc : integer; + var cd : integer; + var ce : integer; + var loop__1__i : integer; + function rmd_hash(message, integer) : chain; + function round_spec(chain, block) : chain; + function rounds(chain, integer, message) : chain; + var ce__1 : integer; + var cd__1 : integer; + var cc__1 : integer; + var cb__1 : integer; + var ca__1 : integer; + +end; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/hash.rls --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/hash.rls Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,61 @@ + /*********************************************************/ + /*Proof Rule Declarations*/ + /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/ + /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/ + /*********************************************************/ + + + /*DATE : 29-NOV-2010 14:30:20.17*/ + + /*function RMD.Hash*/ + + +rule_family hash_rules: + X requires [X:any] & + X <= Y requires [X:ire, Y:ire] & + X >= Y requires [X:ire, Y:ire]. + +hash_rules(1): ca_init may_be_replaced_by 1732584193. +hash_rules(2): cb_init may_be_replaced_by 4023233417. +hash_rules(3): cc_init may_be_replaced_by 2562383102. +hash_rules(4): cd_init may_be_replaced_by 271733878. +hash_rules(5): ce_init may_be_replaced_by 3285377520. +hash_rules(6): interfaces__unsigned_32__size >= 0 may_be_deduced. +hash_rules(7): interfaces__unsigned_32__first may_be_replaced_by 0. +hash_rules(8): interfaces__unsigned_32__last may_be_replaced_by 4294967295. +hash_rules(9): interfaces__unsigned_32__base__first may_be_replaced_by 0. +hash_rules(10): interfaces__unsigned_32__base__last may_be_replaced_by 4294967295. +hash_rules(11): interfaces__unsigned_32__modulus may_be_replaced_by 4294967296. +hash_rules(12): word__size >= 0 may_be_deduced. +hash_rules(13): word__first may_be_replaced_by 0. +hash_rules(14): word__last may_be_replaced_by 4294967295. +hash_rules(15): word__base__first may_be_replaced_by 0. +hash_rules(16): word__base__last may_be_replaced_by 4294967295. +hash_rules(17): word__modulus may_be_replaced_by 4294967296. +hash_rules(18): chain__size >= 0 may_be_deduced. +hash_rules(19): A = B may_be_deduced_from + [goal(checktype(A,chain)), + goal(checktype(B,chain)), + fld_h0(A) = fld_h0(B), + fld_h1(A) = fld_h1(B), + fld_h2(A) = fld_h2(B), + fld_h3(A) = fld_h3(B), + fld_h4(A) = fld_h4(B)]. +hash_rules(20): block_index__size >= 0 may_be_deduced. +hash_rules(21): block_index__first may_be_replaced_by 0. +hash_rules(22): block_index__last may_be_replaced_by 15. +hash_rules(23): block_index__base__first <= block_index__base__last may_be_deduced. +hash_rules(24): block_index__base__first <= block_index__first may_be_deduced. +hash_rules(25): block_index__base__last >= block_index__last may_be_deduced. +hash_rules(26): message_index__size >= 0 may_be_deduced. +hash_rules(27): message_index__first may_be_replaced_by 0. +hash_rules(28): message_index__last may_be_replaced_by 4294967296. +hash_rules(29): message_index__base__first <= message_index__base__last may_be_deduced. +hash_rules(30): message_index__base__first <= message_index__first may_be_deduced. +hash_rules(31): message_index__base__last >= message_index__last may_be_deduced. +hash_rules(32): x__index__subtype__1__first >= message_index__first may_be_deduced. +hash_rules(33): x__index__subtype__1__last <= message_index__last may_be_deduced. +hash_rules(34): x__index__subtype__1__first <= + x__index__subtype__1__last may_be_deduced. +hash_rules(35): x__index__subtype__1__last >= message_index__first may_be_deduced. +hash_rules(36): x__index__subtype__1__first <= message_index__last may_be_deduced. diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/hash.siv --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/hash.siv Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,240 @@ +***************************************************************************** + Semantic Analysis of SPARK Text + Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 + Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. +***************************************************************************** + + +CREATED 29-NOV-2010, 14:30:20 SIMPLIFIED 29-NOV-2010, 14:30:20 + +SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 +Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. + +function RMD.Hash + + + + +For path(s) from start to run-time check associated with statement of line 170: + +function_hash_1. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 171: + +function_hash_2. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 172: + +function_hash_3. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 173: + +function_hash_4. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 174: + +function_hash_5. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 175: + +function_hash_6. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 175: + +function_hash_7. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 177: + +function_hash_8. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 178 to run-time check associated with + statement of line 177: + +function_hash_9. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 177: + +function_hash_10. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 178 to run-time check associated with + statement of line 177: + +function_hash_11. +*** true . /* all conclusions proved */ + + +For path(s) from start to assertion of line 178: + +function_hash_12. +H1: x__index__subtype__1__first = 0 . +H2: for_all(i___2 : integer, 0 <= i___2 and i___2 <= 15 -> for_all(i___1 : + integer, x__index__subtype__1__first <= i___1 and i___1 <= + x__index__subtype__1__last -> 0 <= element(element(x, [i___1]), [ + i___2]) and element(element(x, [i___1]), [i___2]) <= 4294967295)) . +H3: x__index__subtype__1__last >= 0 . +H4: x__index__subtype__1__last <= 4294967296 . +H5: x__index__subtype__1__first <= x__index__subtype__1__last . +H6: mk__chain(h0 := ca__1, h1 := cb__1, h2 := cc__1, h3 := cd__1, h4 := + ce__1) = round_spec(mk__chain(h0 := 1732584193, h1 := 4023233417, h2 + := 2562383102, h3 := 271733878, h4 := 3285377520), element(x, [ + x__index__subtype__1__first])) . +H7: ca__1 >= 0 . +H8: ca__1 <= 4294967295 . +H9: cb__1 >= 0 . +H10: cb__1 <= 4294967295 . +H11: cc__1 >= 0 . +H12: cc__1 <= 4294967295 . +H13: cd__1 >= 0 . +H14: cd__1 <= 4294967295 . +H15: ce__1 >= 0 . +H16: ce__1 <= 4294967295 . + -> +C1: mk__chain(h0 := ca__1, h1 := cb__1, h2 := cc__1, h3 := cd__1, h4 := + ce__1) = rounds(mk__chain(h0 := 1732584193, h1 := 4023233417, h2 := + 2562383102, h3 := 271733878, h4 := 3285377520), + x__index__subtype__1__first + 1, x) . + + +For path(s) from assertion of line 178 to assertion of line 178: + +function_hash_13. +H1: mk__chain(h0 := ca, h1 := cb, h2 := cc, h3 := cd, h4 := ce) = rounds( + mk__chain(h0 := 1732584193, h1 := 4023233417, h2 := 2562383102, h3 := + 271733878, h4 := 3285377520), loop__1__i + 1, x) . +H2: for_all(i___2 : integer, 0 <= i___2 and i___2 <= 15 -> for_all(i___1 : + integer, x__index__subtype__1__first <= i___1 and i___1 <= + x__index__subtype__1__last -> 0 <= element(element(x, [i___1]), [ + i___2]) and element(element(x, [i___1]), [i___2]) <= 4294967295)) . +H3: x__index__subtype__1__first = 0 . +H4: loop__1__i >= 0 . +H5: loop__1__i <= 4294967296 . +H6: loop__1__i >= x__index__subtype__1__first . +H7: ca >= 0 . +H8: ca <= 4294967295 . +H9: cb >= 0 . +H10: cb <= 4294967295 . +H11: cc >= 0 . +H12: cc <= 4294967295 . +H13: cd >= 0 . +H14: cd <= 4294967295 . +H15: ce >= 0 . +H16: ce <= 4294967295 . +H17: loop__1__i + 1 <= x__index__subtype__1__last . +H18: mk__chain(h0 := ca__1, h1 := cb__1, h2 := cc__1, h3 := cd__1, h4 := + ce__1) = round_spec(mk__chain(h0 := ca, h1 := cb, h2 := cc, h3 := cd, + h4 := ce), element(x, [loop__1__i + 1])) . +H19: ca__1 >= 0 . +H20: ca__1 <= 4294967295 . +H21: cb__1 >= 0 . +H22: cb__1 <= 4294967295 . +H23: cc__1 >= 0 . +H24: cc__1 <= 4294967295 . +H25: cd__1 >= 0 . +H26: cd__1 <= 4294967295 . +H27: ce__1 >= 0 . +H28: ce__1 <= 4294967295 . +H29: interfaces__unsigned_32__size >= 0 . +H30: word__size >= 0 . +H31: chain__size >= 0 . +H32: block_index__size >= 0 . +H33: block_index__base__first <= block_index__base__last . +H34: message_index__size >= 0 . +H35: message_index__base__first <= message_index__base__last . +H36: x__index__subtype__1__first <= x__index__subtype__1__last . +H37: block_index__base__first <= 0 . +H38: block_index__base__last >= 15 . +H39: message_index__base__first <= 0 . +H40: x__index__subtype__1__first >= 0 . +H41: x__index__subtype__1__last >= 0 . +H42: message_index__base__last >= 4294967296 . +H43: x__index__subtype__1__last <= 4294967296 . +H44: x__index__subtype__1__first <= 4294967296 . + -> +C1: mk__chain(h0 := ca__1, h1 := cb__1, h2 := cc__1, h3 := cd__1, h4 := + ce__1) = rounds(mk__chain(h0 := 1732584193, h1 := 4023233417, h2 := + 2562383102, h3 := 271733878, h4 := 3285377520), loop__1__i + 2, x) . + + +For path(s) from start to run-time check associated with statement of line 183: + +function_hash_14. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 178 to run-time check associated with + statement of line 183: + +function_hash_15. +*** true . /* all conclusions proved */ + + +For path(s) from start to finish: + +function_hash_16. +*** true . /* contradiction within hypotheses. */ + + + +For path(s) from assertion of line 178 to finish: + +function_hash_17. +H1: mk__chain(h0 := ca, h1 := cb, h2 := cc, h3 := cd, h4 := ce) = rounds( + mk__chain(h0 := 1732584193, h1 := 4023233417, h2 := 2562383102, h3 := + 271733878, h4 := 3285377520), x__index__subtype__1__last + 1, x) . +H2: for_all(i___2 : integer, 0 <= i___2 and i___2 <= 15 -> for_all(i___1 : + integer, x__index__subtype__1__first <= i___1 and i___1 <= + x__index__subtype__1__last -> 0 <= element(element(x, [i___1]), [ + i___2]) and element(element(x, [i___1]), [i___2]) <= 4294967295)) . +H3: x__index__subtype__1__first = 0 . +H4: x__index__subtype__1__last >= 0 . +H5: x__index__subtype__1__last <= 4294967296 . +H6: x__index__subtype__1__last >= x__index__subtype__1__first . +H7: ca >= 0 . +H8: ca <= 4294967295 . +H9: cb >= 0 . +H10: cb <= 4294967295 . +H11: cc >= 0 . +H12: cc <= 4294967295 . +H13: cd >= 0 . +H14: cd <= 4294967295 . +H15: ce >= 0 . +H16: ce <= 4294967295 . +H17: interfaces__unsigned_32__size >= 0 . +H18: word__size >= 0 . +H19: chain__size >= 0 . +H20: block_index__size >= 0 . +H21: block_index__base__first <= block_index__base__last . +H22: message_index__size >= 0 . +H23: message_index__base__first <= message_index__base__last . +H24: x__index__subtype__1__first <= x__index__subtype__1__last . +H25: block_index__base__first <= 0 . +H26: block_index__base__last >= 15 . +H27: message_index__base__first <= 0 . +H28: x__index__subtype__1__first >= 0 . +H29: message_index__base__last >= 4294967296 . +H30: x__index__subtype__1__first <= 4294967296 . + -> +C1: mk__chain(h0 := ca, h1 := cb, h2 := cc, h3 := cd, h4 := ce) = rmd_hash( + x, x__index__subtype__1__last + 1) . + + diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_l.fdl --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_l.fdl Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,38 @@ + {*******************************************************} + {FDL Declarations} + {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039} + {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.} + {*******************************************************} + + + {DATE : 29-NOV-2010 14:30:19.74} + + {function RMD.K_L} + + +title function k_l; + + function round__(real) : integer; + type interfaces__unsigned_32 = integer; + type round_index = integer; + const round_index__base__first : integer = pending; + const round_index__base__last : integer = pending; + const word__base__first : integer = pending; + const word__base__last : integer = pending; + const interfaces__unsigned_32__base__first : integer = pending; + const interfaces__unsigned_32__base__last : integer = pending; + const round_index__first : integer = pending; + const round_index__last : integer = pending; + const round_index__size : integer = pending; + const word__first : integer = pending; + const word__last : integer = pending; + const word__modulus : integer = pending; + const word__size : integer = pending; + const interfaces__unsigned_32__first : integer = pending; + const interfaces__unsigned_32__last : integer = pending; + const interfaces__unsigned_32__modulus : integer = pending; + const interfaces__unsigned_32__size : integer = pending; + var j : integer; + function k_l_spec(integer) : integer; + +end; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_l.rls --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_l.rls Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,35 @@ + /*********************************************************/ + /*Proof Rule Declarations*/ + /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/ + /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/ + /*********************************************************/ + + + /*DATE : 29-NOV-2010 14:30:19.74*/ + + /*function RMD.K_L*/ + + +rule_family k_l_rules: + X requires [X:any] & + X <= Y requires [X:ire, Y:ire] & + X >= Y requires [X:ire, Y:ire]. + +k_l_rules(1): interfaces__unsigned_32__size >= 0 may_be_deduced. +k_l_rules(2): interfaces__unsigned_32__first may_be_replaced_by 0. +k_l_rules(3): interfaces__unsigned_32__last may_be_replaced_by 4294967295. +k_l_rules(4): interfaces__unsigned_32__base__first may_be_replaced_by 0. +k_l_rules(5): interfaces__unsigned_32__base__last may_be_replaced_by 4294967295. +k_l_rules(6): interfaces__unsigned_32__modulus may_be_replaced_by 4294967296. +k_l_rules(7): word__size >= 0 may_be_deduced. +k_l_rules(8): word__first may_be_replaced_by 0. +k_l_rules(9): word__last may_be_replaced_by 4294967295. +k_l_rules(10): word__base__first may_be_replaced_by 0. +k_l_rules(11): word__base__last may_be_replaced_by 4294967295. +k_l_rules(12): word__modulus may_be_replaced_by 4294967296. +k_l_rules(13): round_index__size >= 0 may_be_deduced. +k_l_rules(14): round_index__first may_be_replaced_by 0. +k_l_rules(15): round_index__last may_be_replaced_by 79. +k_l_rules(16): round_index__base__first <= round_index__base__last may_be_deduced. +k_l_rules(17): round_index__base__first <= round_index__first may_be_deduced. +k_l_rules(18): round_index__base__last >= round_index__last may_be_deduced. diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_l.siv --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_l.siv Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,118 @@ +***************************************************************************** + Semantic Analysis of SPARK Text + Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 + Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. +***************************************************************************** + + +CREATED 29-NOV-2010, 14:30:19 SIMPLIFIED 29-NOV-2010, 14:30:28 + +SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 +Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. + +function RMD.K_L + + + + +For path(s) from start to run-time check associated with statement of line 24: + +function_k_l_1. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 25: + +function_k_l_2. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 26: + +function_k_l_3. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 27: + +function_k_l_4. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 28: + +function_k_l_5. +*** true . /* all conclusions proved */ + + +For path(s) from start to finish: + +function_k_l_6. +H1: j >= 0 . +H2: j <= 15 . +H3: interfaces__unsigned_32__size >= 0 . +H4: word__size >= 0 . +H5: round_index__size >= 0 . +H6: round_index__base__first <= round_index__base__last . +H7: round_index__base__first <= 0 . +H8: round_index__base__last >= 79 . + -> +C1: 0 = k_l_spec(j) . + + +function_k_l_7. +H1: 16 <= j . +H2: j <= 31 . +H3: interfaces__unsigned_32__size >= 0 . +H4: word__size >= 0 . +H5: round_index__size >= 0 . +H6: round_index__base__first <= round_index__base__last . +H7: round_index__base__first <= 0 . +H8: round_index__base__last >= 79 . + -> +C1: 1518500249 = k_l_spec(j) . + + +function_k_l_8. +H1: 32 <= j . +H2: j <= 47 . +H3: interfaces__unsigned_32__size >= 0 . +H4: word__size >= 0 . +H5: round_index__size >= 0 . +H6: round_index__base__first <= round_index__base__last . +H7: round_index__base__first <= 0 . +H8: round_index__base__last >= 79 . + -> +C1: 1859775393 = k_l_spec(j) . + + +function_k_l_9. +H1: 48 <= j . +H2: j <= 63 . +H3: interfaces__unsigned_32__size >= 0 . +H4: word__size >= 0 . +H5: round_index__size >= 0 . +H6: round_index__base__first <= round_index__base__last . +H7: round_index__base__first <= 0 . +H8: round_index__base__last >= 79 . + -> +C1: 2400959708 = k_l_spec(j) . + + +function_k_l_10. +H1: j >= 0 . +H2: j <= 79 . +H3: 15 < j . +H4: 31 < j . +H5: 47 < j . +H6: 63 < j . +H7: interfaces__unsigned_32__size >= 0 . +H8: word__size >= 0 . +H9: round_index__size >= 0 . +H10: round_index__base__first <= round_index__base__last . +H11: round_index__base__first <= 0 . +H12: round_index__base__last >= 79 . + -> +C1: 2840853838 = k_l_spec(j) . + + diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_r.fdl --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_r.fdl Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,38 @@ + {*******************************************************} + {FDL Declarations} + {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039} + {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.} + {*******************************************************} + + + {DATE : 29-NOV-2010 14:30:19.76} + + {function RMD.K_R} + + +title function k_r; + + function round__(real) : integer; + type interfaces__unsigned_32 = integer; + type round_index = integer; + const round_index__base__first : integer = pending; + const round_index__base__last : integer = pending; + const word__base__first : integer = pending; + const word__base__last : integer = pending; + const interfaces__unsigned_32__base__first : integer = pending; + const interfaces__unsigned_32__base__last : integer = pending; + const round_index__first : integer = pending; + const round_index__last : integer = pending; + const round_index__size : integer = pending; + const word__first : integer = pending; + const word__last : integer = pending; + const word__modulus : integer = pending; + const word__size : integer = pending; + const interfaces__unsigned_32__first : integer = pending; + const interfaces__unsigned_32__last : integer = pending; + const interfaces__unsigned_32__modulus : integer = pending; + const interfaces__unsigned_32__size : integer = pending; + var j : integer; + function k_r_spec(integer) : integer; + +end; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_r.rls --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_r.rls Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,35 @@ + /*********************************************************/ + /*Proof Rule Declarations*/ + /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/ + /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/ + /*********************************************************/ + + + /*DATE : 29-NOV-2010 14:30:19.76*/ + + /*function RMD.K_R*/ + + +rule_family k_r_rules: + X requires [X:any] & + X <= Y requires [X:ire, Y:ire] & + X >= Y requires [X:ire, Y:ire]. + +k_r_rules(1): interfaces__unsigned_32__size >= 0 may_be_deduced. +k_r_rules(2): interfaces__unsigned_32__first may_be_replaced_by 0. +k_r_rules(3): interfaces__unsigned_32__last may_be_replaced_by 4294967295. +k_r_rules(4): interfaces__unsigned_32__base__first may_be_replaced_by 0. +k_r_rules(5): interfaces__unsigned_32__base__last may_be_replaced_by 4294967295. +k_r_rules(6): interfaces__unsigned_32__modulus may_be_replaced_by 4294967296. +k_r_rules(7): word__size >= 0 may_be_deduced. +k_r_rules(8): word__first may_be_replaced_by 0. +k_r_rules(9): word__last may_be_replaced_by 4294967295. +k_r_rules(10): word__base__first may_be_replaced_by 0. +k_r_rules(11): word__base__last may_be_replaced_by 4294967295. +k_r_rules(12): word__modulus may_be_replaced_by 4294967296. +k_r_rules(13): round_index__size >= 0 may_be_deduced. +k_r_rules(14): round_index__first may_be_replaced_by 0. +k_r_rules(15): round_index__last may_be_replaced_by 79. +k_r_rules(16): round_index__base__first <= round_index__base__last may_be_deduced. +k_r_rules(17): round_index__base__first <= round_index__first may_be_deduced. +k_r_rules(18): round_index__base__last >= round_index__last may_be_deduced. diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_r.siv --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_r.siv Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,118 @@ +***************************************************************************** + Semantic Analysis of SPARK Text + Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 + Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. +***************************************************************************** + + +CREATED 29-NOV-2010, 14:30:19 SIMPLIFIED 29-NOV-2010, 14:30:21 + +SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 +Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. + +function RMD.K_R + + + + +For path(s) from start to run-time check associated with statement of line 38: + +function_k_r_1. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 39: + +function_k_r_2. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 40: + +function_k_r_3. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 41: + +function_k_r_4. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 42: + +function_k_r_5. +*** true . /* all conclusions proved */ + + +For path(s) from start to finish: + +function_k_r_6. +H1: j >= 0 . +H2: j <= 15 . +H3: interfaces__unsigned_32__size >= 0 . +H4: word__size >= 0 . +H5: round_index__size >= 0 . +H6: round_index__base__first <= round_index__base__last . +H7: round_index__base__first <= 0 . +H8: round_index__base__last >= 79 . + -> +C1: 1352829926 = k_r_spec(j) . + + +function_k_r_7. +H1: 16 <= j . +H2: j <= 31 . +H3: interfaces__unsigned_32__size >= 0 . +H4: word__size >= 0 . +H5: round_index__size >= 0 . +H6: round_index__base__first <= round_index__base__last . +H7: round_index__base__first <= 0 . +H8: round_index__base__last >= 79 . + -> +C1: 1548603684 = k_r_spec(j) . + + +function_k_r_8. +H1: 32 <= j . +H2: j <= 47 . +H3: interfaces__unsigned_32__size >= 0 . +H4: word__size >= 0 . +H5: round_index__size >= 0 . +H6: round_index__base__first <= round_index__base__last . +H7: round_index__base__first <= 0 . +H8: round_index__base__last >= 79 . + -> +C1: 1836072691 = k_r_spec(j) . + + +function_k_r_9. +H1: 48 <= j . +H2: j <= 63 . +H3: interfaces__unsigned_32__size >= 0 . +H4: word__size >= 0 . +H5: round_index__size >= 0 . +H6: round_index__base__first <= round_index__base__last . +H7: round_index__base__first <= 0 . +H8: round_index__base__last >= 79 . + -> +C1: 2053994217 = k_r_spec(j) . + + +function_k_r_10. +H1: j >= 0 . +H2: j <= 79 . +H3: 15 < j . +H4: 31 < j . +H5: 47 < j . +H6: 63 < j . +H7: interfaces__unsigned_32__size >= 0 . +H8: word__size >= 0 . +H9: round_index__size >= 0 . +H10: round_index__base__first <= round_index__base__last . +H11: round_index__base__first <= 0 . +H12: round_index__base__last >= 79 . + -> +C1: 0 = k_r_spec(j) . + + diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_l.fdl --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_l.fdl Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,33 @@ + {*******************************************************} + {FDL Declarations} + {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039} + {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.} + {*******************************************************} + + + {DATE : 29-NOV-2010 14:30:19.77} + + {function RMD.R_L} + + +title function r_l; + + function round__(real) : integer; + type block_index = integer; + type round_index = integer; + type block_permutation = array [integer] of integer; + const r_values : block_permutation = pending; + const round_index__base__first : integer = pending; + const round_index__base__last : integer = pending; + const block_index__base__first : integer = pending; + const block_index__base__last : integer = pending; + const round_index__first : integer = pending; + const round_index__last : integer = pending; + const round_index__size : integer = pending; + const block_index__first : integer = pending; + const block_index__last : integer = pending; + const block_index__size : integer = pending; + var j : integer; + function r_l_spec(integer) : integer; + +end; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_l.rls --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_l.rls Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,75 @@ + /*********************************************************/ + /*Proof Rule Declarations*/ + /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/ + /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/ + /*********************************************************/ + + + /*DATE : 29-NOV-2010 14:30:19.77*/ + + /*function RMD.R_L*/ + + +rule_family r_l_rules: + X requires [X:any] & + X <= Y requires [X:ire, Y:ire] & + X >= Y requires [X:ire, Y:ire]. + +r_l_rules(1): block_index__first <= element(r_values, [I]) may_be_deduced_from [0 <= I, I <= 79]. +r_l_rules(2): element(r_values, [I]) <= block_index__last may_be_deduced_from [0 <= I, I <= 79]. +r_l_rules(3): r_values may_be_replaced_by + mk__block_permutation([round_index__first] := 0, [ + round_index__first + 1] := 1, [round_index__first + 2] := 2, [ + round_index__first + 3] := 3, [round_index__first + 4] := 4, [ + round_index__first + 5] := 5, [round_index__first + 6] := 6, [ + round_index__first + 7] := 7, [round_index__first + 8] := 8, [ + round_index__first + 9] := 9, [round_index__first + 10] := + 10, [round_index__first + 11] := 11, [ + round_index__first + 12] := 12, [round_index__first + 13] := + 13, [round_index__first + 14] := 14, [ + round_index__first + 15] := 15, [round_index__first + 16] := + 7, [round_index__first + 17] := 4, [round_index__first + 18] := + 13, [round_index__first + 19] := 1, [round_index__first + 20] := + 10, [round_index__first + 21] := 6, [round_index__first + 22] := + 15, [round_index__first + 23] := 3, [round_index__first + 24] := + 12, [round_index__first + 25] := 0, [round_index__first + 26] := + 9, [round_index__first + 27] := 5, [round_index__first + 28] := + 2, [round_index__first + 29] := 14, [round_index__first + 30] := + 11, [round_index__first + 31] := 8, [round_index__first + 32] := + 3, [round_index__first + 33] := 10, [round_index__first + 34] := + 14, [round_index__first + 35] := 4, [round_index__first + 36] := + 9, [round_index__first + 37] := 15, [round_index__first + 38] := + 8, [round_index__first + 39] := 1, [round_index__first + 40] := + 2, [round_index__first + 41] := 7, [round_index__first + 42] := + 0, [round_index__first + 43] := 6, [round_index__first + 44] := + 13, [round_index__first + 45] := 11, [ + round_index__first + 46] := 5, [round_index__first + 47] := + 12, [round_index__first + 48] := 1, [round_index__first + 49] := + 9, [round_index__first + 50] := 11, [round_index__first + 51] := + 10, [round_index__first + 52] := 0, [round_index__first + 53] := + 8, [round_index__first + 54] := 12, [round_index__first + 55] := + 4, [round_index__first + 56] := 13, [round_index__first + 57] := + 3, [round_index__first + 58] := 7, [round_index__first + 59] := + 15, [round_index__first + 60] := 14, [ + round_index__first + 61] := 5, [round_index__first + 62] := + 6, [round_index__first + 63] := 2, [round_index__first + 64] := + 4, [round_index__first + 65] := 0, [round_index__first + 66] := + 5, [round_index__first + 67] := 9, [round_index__first + 68] := + 7, [round_index__first + 69] := 12, [round_index__first + 70] := + 2, [round_index__first + 71] := 10, [round_index__first + 72] := + 14, [round_index__first + 73] := 1, [round_index__first + 74] := + 3, [round_index__first + 75] := 8, [round_index__first + 76] := + 11, [round_index__first + 77] := 6, [round_index__first + 78] := + 15, [round_index__first + 79] := 13). +r_l_rules(4): block_index__size >= 0 may_be_deduced. +r_l_rules(5): block_index__first may_be_replaced_by 0. +r_l_rules(6): block_index__last may_be_replaced_by 15. +r_l_rules(7): block_index__base__first <= block_index__base__last may_be_deduced. +r_l_rules(8): block_index__base__first <= block_index__first may_be_deduced. +r_l_rules(9): block_index__base__last >= block_index__last may_be_deduced. +r_l_rules(10): round_index__size >= 0 may_be_deduced. +r_l_rules(11): round_index__first may_be_replaced_by 0. +r_l_rules(12): round_index__last may_be_replaced_by 79. +r_l_rules(13): round_index__base__first <= round_index__base__last may_be_deduced. +r_l_rules(14): round_index__base__first <= round_index__first may_be_deduced. +r_l_rules(15): round_index__base__last >= round_index__last may_be_deduced. diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_l.siv --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_l.siv Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,81 @@ +***************************************************************************** + Semantic Analysis of SPARK Text + Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 + Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. +***************************************************************************** + + +CREATED 29-NOV-2010, 14:30:19 SIMPLIFIED 29-NOV-2010, 14:30:28 + +SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 +Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. + +function RMD.R_L + + + + +For path(s) from start to run-time check associated with statement of line 59: + +function_r_l_1. +*** true . /* all conclusions proved */ + + +For path(s) from start to finish: + +function_r_l_2. +H1: j >= 0 . +H2: j <= 79 . +H3: block_index__size >= 0 . +H4: block_index__base__first <= block_index__base__last . +H5: round_index__size >= 0 . +H6: round_index__base__first <= round_index__base__last . +H7: block_index__base__first <= 0 . +H8: block_index__base__last >= 15 . +H9: round_index__base__first <= 0 . +H10: round_index__base__last >= 79 . + -> +C1: element(mk__block_permutation([0] := 0, [1] := 1, [2] := 2, [3] := 3, [4] + := 4, [5] := 5, [6] := 6, [7] := 7, [8] := 8, [9] := 9, [10] := 10, [ + 11] := 11, [12] := 12, [13] := 13, [14] := 14, [15] := 15, [16] := 7, + [17] := 4, [18] := 13, [19] := 1, [20] := 10, [21] := 6, [22] := 15, [ + 23] := 3, [24] := 12, [25] := 0, [26] := 9, [27] := 5, [28] := 2, [29] + := 14, [30] := 11, [31] := 8, [32] := 3, [33] := 10, [34] := 14, [35] + := 4, [36] := 9, [37] := 15, [38] := 8, [39] := 1, [40] := 2, [41] + := 7, [42] := 0, [43] := 6, [44] := 13, [45] := 11, [46] := 5, [47] + := 12, [48] := 1, [49] := 9, [50] := 11, [51] := 10, [52] := 0, [53] + := 8, [54] := 12, [55] := 4, [56] := 13, [57] := 3, [58] := 7, [59] + := 15, [60] := 14, [61] := 5, [62] := 6, [63] := 2, [64] := 4, [65] + := 0, [66] := 5, [67] := 9, [68] := 7, [69] := 12, [70] := 2, [71] := + 10, [72] := 14, [73] := 1, [74] := 3, [75] := 8, [76] := 11, [77] := + 6, [78] := 15, [79] := 13), [j]) = r_l_spec(j) . +C2: element(mk__block_permutation([0] := 0, [1] := 1, [2] := 2, [3] := 3, [4] + := 4, [5] := 5, [6] := 6, [7] := 7, [8] := 8, [9] := 9, [10] := 10, [ + 11] := 11, [12] := 12, [13] := 13, [14] := 14, [15] := 15, [16] := 7, + [17] := 4, [18] := 13, [19] := 1, [20] := 10, [21] := 6, [22] := 15, [ + 23] := 3, [24] := 12, [25] := 0, [26] := 9, [27] := 5, [28] := 2, [29] + := 14, [30] := 11, [31] := 8, [32] := 3, [33] := 10, [34] := 14, [35] + := 4, [36] := 9, [37] := 15, [38] := 8, [39] := 1, [40] := 2, [41] + := 7, [42] := 0, [43] := 6, [44] := 13, [45] := 11, [46] := 5, [47] + := 12, [48] := 1, [49] := 9, [50] := 11, [51] := 10, [52] := 0, [53] + := 8, [54] := 12, [55] := 4, [56] := 13, [57] := 3, [58] := 7, [59] + := 15, [60] := 14, [61] := 5, [62] := 6, [63] := 2, [64] := 4, [65] + := 0, [66] := 5, [67] := 9, [68] := 7, [69] := 12, [70] := 2, [71] := + 10, [72] := 14, [73] := 1, [74] := 3, [75] := 8, [76] := 11, [77] := + 6, [78] := 15, [79] := 13), [j]) >= 0 . +C3: element(mk__block_permutation([0] := 0, [1] := 1, [2] := 2, [3] := 3, [4] + := 4, [5] := 5, [6] := 6, [7] := 7, [8] := 8, [9] := 9, [10] := 10, [ + 11] := 11, [12] := 12, [13] := 13, [14] := 14, [15] := 15, [16] := 7, + [17] := 4, [18] := 13, [19] := 1, [20] := 10, [21] := 6, [22] := 15, [ + 23] := 3, [24] := 12, [25] := 0, [26] := 9, [27] := 5, [28] := 2, [29] + := 14, [30] := 11, [31] := 8, [32] := 3, [33] := 10, [34] := 14, [35] + := 4, [36] := 9, [37] := 15, [38] := 8, [39] := 1, [40] := 2, [41] + := 7, [42] := 0, [43] := 6, [44] := 13, [45] := 11, [46] := 5, [47] + := 12, [48] := 1, [49] := 9, [50] := 11, [51] := 10, [52] := 0, [53] + := 8, [54] := 12, [55] := 4, [56] := 13, [57] := 3, [58] := 7, [59] + := 15, [60] := 14, [61] := 5, [62] := 6, [63] := 2, [64] := 4, [65] + := 0, [66] := 5, [67] := 9, [68] := 7, [69] := 12, [70] := 2, [71] := + 10, [72] := 14, [73] := 1, [74] := 3, [75] := 8, [76] := 11, [77] := + 6, [78] := 15, [79] := 13), [j]) <= 15 . + + diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_r.fdl --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_r.fdl Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,33 @@ + {*******************************************************} + {FDL Declarations} + {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039} + {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.} + {*******************************************************} + + + {DATE : 29-NOV-2010 14:30:19.81} + + {function RMD.R_R} + + +title function r_r; + + function round__(real) : integer; + type block_index = integer; + type round_index = integer; + type block_permutation = array [integer] of integer; + const r_values : block_permutation = pending; + const round_index__base__first : integer = pending; + const round_index__base__last : integer = pending; + const block_index__base__first : integer = pending; + const block_index__base__last : integer = pending; + const round_index__first : integer = pending; + const round_index__last : integer = pending; + const round_index__size : integer = pending; + const block_index__first : integer = pending; + const block_index__last : integer = pending; + const block_index__size : integer = pending; + var j : integer; + function r_r_spec(integer) : integer; + +end; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_r.rls --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_r.rls Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,76 @@ + /*********************************************************/ + /*Proof Rule Declarations*/ + /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/ + /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/ + /*********************************************************/ + + + /*DATE : 29-NOV-2010 14:30:19.81*/ + + /*function RMD.R_R*/ + + +rule_family r_r_rules: + X requires [X:any] & + X <= Y requires [X:ire, Y:ire] & + X >= Y requires [X:ire, Y:ire]. + +r_r_rules(1): block_index__first <= element(r_values, [I]) may_be_deduced_from [0 <= I, I <= 79]. +r_r_rules(2): element(r_values, [I]) <= block_index__last may_be_deduced_from [0 <= I, I <= 79]. +r_r_rules(3): r_values may_be_replaced_by + mk__block_permutation([round_index__first] := 5, [ + round_index__first + 1] := 14, [round_index__first + 2] := 7, [ + round_index__first + 3] := 0, [round_index__first + 4] := 9, [ + round_index__first + 5] := 2, [round_index__first + 6] := 11, [ + round_index__first + 7] := 4, [round_index__first + 8] := 13, [ + round_index__first + 9] := 6, [round_index__first + 10] := + 15, [round_index__first + 11] := 8, [round_index__first + 12] := + 1, [round_index__first + 13] := 10, [round_index__first + 14] := + 3, [round_index__first + 15] := 12, [round_index__first + 16] := + 6, [round_index__first + 17] := 11, [round_index__first + 18] := + 3, [round_index__first + 19] := 7, [round_index__first + 20] := + 0, [round_index__first + 21] := 13, [round_index__first + 22] := + 5, [round_index__first + 23] := 10, [round_index__first + 24] := + 14, [round_index__first + 25] := 15, [ + round_index__first + 26] := 8, [round_index__first + 27] := + 12, [round_index__first + 28] := 4, [round_index__first + 29] := + 9, [round_index__first + 30] := 1, [round_index__first + 31] := + 2, [round_index__first + 32] := 15, [round_index__first + 33] := + 5, [round_index__first + 34] := 1, [round_index__first + 35] := + 3, [round_index__first + 36] := 7, [round_index__first + 37] := + 14, [round_index__first + 38] := 6, [round_index__first + 39] := + 9, [round_index__first + 40] := 11, [round_index__first + 41] := + 8, [round_index__first + 42] := 12, [round_index__first + 43] := + 2, [round_index__first + 44] := 10, [round_index__first + 45] := + 0, [round_index__first + 46] := 4, [round_index__first + 47] := + 13, [round_index__first + 48] := 8, [round_index__first + 49] := + 6, [round_index__first + 50] := 4, [round_index__first + 51] := + 1, [round_index__first + 52] := 3, [round_index__first + 53] := + 11, [round_index__first + 54] := 15, [ + round_index__first + 55] := 0, [round_index__first + 56] := + 5, [round_index__first + 57] := 12, [round_index__first + 58] := + 2, [round_index__first + 59] := 13, [round_index__first + 60] := + 9, [round_index__first + 61] := 7, [round_index__first + 62] := + 10, [round_index__first + 63] := 14, [ + round_index__first + 64] := 12, [round_index__first + 65] := + 15, [round_index__first + 66] := 10, [ + round_index__first + 67] := 4, [round_index__first + 68] := + 1, [round_index__first + 69] := 5, [round_index__first + 70] := + 8, [round_index__first + 71] := 7, [round_index__first + 72] := + 6, [round_index__first + 73] := 2, [round_index__first + 74] := + 13, [round_index__first + 75] := 14, [ + round_index__first + 76] := 0, [round_index__first + 77] := + 3, [round_index__first + 78] := 9, [round_index__first + 79] := + 11). +r_r_rules(4): block_index__size >= 0 may_be_deduced. +r_r_rules(5): block_index__first may_be_replaced_by 0. +r_r_rules(6): block_index__last may_be_replaced_by 15. +r_r_rules(7): block_index__base__first <= block_index__base__last may_be_deduced. +r_r_rules(8): block_index__base__first <= block_index__first may_be_deduced. +r_r_rules(9): block_index__base__last >= block_index__last may_be_deduced. +r_r_rules(10): round_index__size >= 0 may_be_deduced. +r_r_rules(11): round_index__first may_be_replaced_by 0. +r_r_rules(12): round_index__last may_be_replaced_by 79. +r_r_rules(13): round_index__base__first <= round_index__base__last may_be_deduced. +r_r_rules(14): round_index__base__first <= round_index__first may_be_deduced. +r_r_rules(15): round_index__base__last >= round_index__last may_be_deduced. diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_r.siv --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_r.siv Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,81 @@ +***************************************************************************** + Semantic Analysis of SPARK Text + Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 + Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. +***************************************************************************** + + +CREATED 29-NOV-2010, 14:30:19 SIMPLIFIED 29-NOV-2010, 14:30:21 + +SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 +Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. + +function RMD.R_R + + + + +For path(s) from start to run-time check associated with statement of line 73: + +function_r_r_1. +*** true . /* all conclusions proved */ + + +For path(s) from start to finish: + +function_r_r_2. +H1: j >= 0 . +H2: j <= 79 . +H3: block_index__size >= 0 . +H4: block_index__base__first <= block_index__base__last . +H5: round_index__size >= 0 . +H6: round_index__base__first <= round_index__base__last . +H7: block_index__base__first <= 0 . +H8: block_index__base__last >= 15 . +H9: round_index__base__first <= 0 . +H10: round_index__base__last >= 79 . + -> +C1: element(mk__block_permutation([0] := 5, [1] := 14, [2] := 7, [3] := 0, [ + 4] := 9, [5] := 2, [6] := 11, [7] := 4, [8] := 13, [9] := 6, [10] := + 15, [11] := 8, [12] := 1, [13] := 10, [14] := 3, [15] := 12, [16] := + 6, [17] := 11, [18] := 3, [19] := 7, [20] := 0, [21] := 13, [22] := + 5, [23] := 10, [24] := 14, [25] := 15, [26] := 8, [27] := 12, [28] := + 4, [29] := 9, [30] := 1, [31] := 2, [32] := 15, [33] := 5, [34] := 1, + [35] := 3, [36] := 7, [37] := 14, [38] := 6, [39] := 9, [40] := 11, [ + 41] := 8, [42] := 12, [43] := 2, [44] := 10, [45] := 0, [46] := 4, [ + 47] := 13, [48] := 8, [49] := 6, [50] := 4, [51] := 1, [52] := 3, [53] + := 11, [54] := 15, [55] := 0, [56] := 5, [57] := 12, [58] := 2, [59] + := 13, [60] := 9, [61] := 7, [62] := 10, [63] := 14, [64] := 12, [65] + := 15, [66] := 10, [67] := 4, [68] := 1, [69] := 5, [70] := 8, [71] + := 7, [72] := 6, [73] := 2, [74] := 13, [75] := 14, [76] := 0, [77] + := 3, [78] := 9, [79] := 11), [j]) = r_r_spec(j) . +C2: element(mk__block_permutation([0] := 5, [1] := 14, [2] := 7, [3] := 0, [ + 4] := 9, [5] := 2, [6] := 11, [7] := 4, [8] := 13, [9] := 6, [10] := + 15, [11] := 8, [12] := 1, [13] := 10, [14] := 3, [15] := 12, [16] := + 6, [17] := 11, [18] := 3, [19] := 7, [20] := 0, [21] := 13, [22] := + 5, [23] := 10, [24] := 14, [25] := 15, [26] := 8, [27] := 12, [28] := + 4, [29] := 9, [30] := 1, [31] := 2, [32] := 15, [33] := 5, [34] := 1, + [35] := 3, [36] := 7, [37] := 14, [38] := 6, [39] := 9, [40] := 11, [ + 41] := 8, [42] := 12, [43] := 2, [44] := 10, [45] := 0, [46] := 4, [ + 47] := 13, [48] := 8, [49] := 6, [50] := 4, [51] := 1, [52] := 3, [53] + := 11, [54] := 15, [55] := 0, [56] := 5, [57] := 12, [58] := 2, [59] + := 13, [60] := 9, [61] := 7, [62] := 10, [63] := 14, [64] := 12, [65] + := 15, [66] := 10, [67] := 4, [68] := 1, [69] := 5, [70] := 8, [71] + := 7, [72] := 6, [73] := 2, [74] := 13, [75] := 14, [76] := 0, [77] + := 3, [78] := 9, [79] := 11), [j]) >= 0 . +C3: element(mk__block_permutation([0] := 5, [1] := 14, [2] := 7, [3] := 0, [ + 4] := 9, [5] := 2, [6] := 11, [7] := 4, [8] := 13, [9] := 6, [10] := + 15, [11] := 8, [12] := 1, [13] := 10, [14] := 3, [15] := 12, [16] := + 6, [17] := 11, [18] := 3, [19] := 7, [20] := 0, [21] := 13, [22] := + 5, [23] := 10, [24] := 14, [25] := 15, [26] := 8, [27] := 12, [28] := + 4, [29] := 9, [30] := 1, [31] := 2, [32] := 15, [33] := 5, [34] := 1, + [35] := 3, [36] := 7, [37] := 14, [38] := 6, [39] := 9, [40] := 11, [ + 41] := 8, [42] := 12, [43] := 2, [44] := 10, [45] := 0, [46] := 4, [ + 47] := 13, [48] := 8, [49] := 6, [50] := 4, [51] := 1, [52] := 3, [53] + := 11, [54] := 15, [55] := 0, [56] := 5, [57] := 12, [58] := 2, [59] + := 13, [60] := 9, [61] := 7, [62] := 10, [63] := 14, [64] := 12, [65] + := 15, [66] := 10, [67] := 4, [68] := 1, [69] := 5, [70] := 8, [71] + := 7, [72] := 6, [73] := 2, [74] := 13, [75] := 14, [76] := 0, [77] + := 3, [78] := 9, [79] := 11), [j]) <= 15 . + + diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/round.fdl --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/round.fdl Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,112 @@ + {*******************************************************} + {FDL Declarations} + {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039} + {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.} + {*******************************************************} + + + {DATE : 29-NOV-2010 14:30:19.87} + + {procedure RMD.Round} + + +title procedure round; + + function round__(real) : integer; + type interfaces__unsigned_32 = integer; + type block_index = integer; + type round_index = integer; + type chain = record + h0 : integer; + h1 : integer; + h2 : integer; + h3 : integer; + h4 : integer + end; + type block = array [integer] of integer; + type chain_pair = record + left : chain; + right : chain + end; + const rotate_amount__base__first : integer = pending; + const rotate_amount__base__last : integer = pending; + const round_index__base__first : integer = pending; + const round_index__base__last : integer = pending; + const block_index__base__first : integer = pending; + const block_index__base__last : integer = pending; + const word__base__first : integer = pending; + const word__base__last : integer = pending; + const wordops__rotate_amount__base__first : integer = pending; + const wordops__rotate_amount__base__last : integer = pending; + const wordops__word__base__first : integer = pending; + const wordops__word__base__last : integer = pending; + const interfaces__unsigned_32__base__first : integer = pending; + const interfaces__unsigned_32__base__last : integer = pending; + const integer__base__first : integer = pending; + const integer__base__last : integer = pending; + const rotate_amount__first : integer = pending; + const rotate_amount__last : integer = pending; + const rotate_amount__size : integer = pending; + const chain_pair__size : integer = pending; + const round_index__first : integer = pending; + const round_index__last : integer = pending; + const round_index__size : integer = pending; + const block_index__first : integer = pending; + const block_index__last : integer = pending; + const block_index__size : integer = pending; + const chain__size : integer = pending; + const word__first : integer = pending; + const word__last : integer = pending; + const word__modulus : integer = pending; + const word__size : integer = pending; + const wordops__rotate_amount__first : integer = pending; + const wordops__rotate_amount__last : integer = pending; + const wordops__rotate_amount__size : integer = pending; + const wordops__word__first : integer = pending; + const wordops__word__last : integer = pending; + const wordops__word__modulus : integer = pending; + const wordops__word__size : integer = pending; + const interfaces__unsigned_32__first : integer = pending; + const interfaces__unsigned_32__last : integer = pending; + const interfaces__unsigned_32__modulus : integer = pending; + const interfaces__unsigned_32__size : integer = pending; + const integer__first : integer = pending; + const integer__last : integer = pending; + const integer__size : integer = pending; + var ca : integer; + var cb : integer; + var cc : integer; + var cd : integer; + var ce : integer; + var x : block; + var cla : integer; + var clb : integer; + var clc : integer; + var cld : integer; + var cle : integer; + var cra : integer; + var crb : integer; + var crc : integer; + var crd : integer; + var cre : integer; + var loop__1__j : integer; + function wordops__rotate_left(integer, integer) : integer; + function wordops__rotate(integer, integer) : integer; + function f_spec(integer, integer, integer, integer) : integer; + function k_l_spec(integer) : integer; + function k_r_spec(integer) : integer; + function r_l_spec(integer) : integer; + function r_r_spec(integer) : integer; + function s_l_spec(integer) : integer; + function s_r_spec(integer) : integer; + function steps(chain_pair, integer, block) : chain_pair; + function round_spec(chain, block) : chain; + function f(integer, integer, integer, integer) : integer; + function k_l(integer) : integer; + function k_r(integer) : integer; + function r_l(integer) : integer; + function r_r(integer) : integer; + function s_l(integer) : integer; + function s_r(integer) : integer; + +end; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/round.rls --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/round.rls Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,77 @@ + /*********************************************************/ + /*Proof Rule Declarations*/ + /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/ + /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/ + /*********************************************************/ + + + /*DATE : 29-NOV-2010 14:30:19.87*/ + + /*procedure RMD.Round*/ + + +rule_family round_rules: + X requires [X:any] & + X <= Y requires [X:ire, Y:ire] & + X >= Y requires [X:ire, Y:ire]. + +round_rules(1): integer__size >= 0 may_be_deduced. +round_rules(2): integer__first may_be_replaced_by -2147483648. +round_rules(3): integer__last may_be_replaced_by 2147483647. +round_rules(4): integer__base__first may_be_replaced_by -2147483648. +round_rules(5): integer__base__last may_be_replaced_by 2147483647. +round_rules(6): interfaces__unsigned_32__size >= 0 may_be_deduced. +round_rules(7): interfaces__unsigned_32__first may_be_replaced_by 0. +round_rules(8): interfaces__unsigned_32__last may_be_replaced_by 4294967295. +round_rules(9): interfaces__unsigned_32__base__first may_be_replaced_by 0. +round_rules(10): interfaces__unsigned_32__base__last may_be_replaced_by 4294967295. +round_rules(11): interfaces__unsigned_32__modulus may_be_replaced_by 4294967296. +round_rules(12): wordops__word__size >= 0 may_be_deduced. +round_rules(13): wordops__word__first may_be_replaced_by 0. +round_rules(14): wordops__word__last may_be_replaced_by 4294967295. +round_rules(15): wordops__word__base__first may_be_replaced_by 0. +round_rules(16): wordops__word__base__last may_be_replaced_by 4294967295. +round_rules(17): wordops__word__modulus may_be_replaced_by 4294967296. +round_rules(18): wordops__rotate_amount__size >= 0 may_be_deduced. +round_rules(19): wordops__rotate_amount__first may_be_replaced_by 0. +round_rules(20): wordops__rotate_amount__last may_be_replaced_by 15. +round_rules(21): wordops__rotate_amount__base__first may_be_replaced_by -2147483648. +round_rules(22): wordops__rotate_amount__base__last may_be_replaced_by 2147483647. +round_rules(23): word__size >= 0 may_be_deduced. +round_rules(24): word__first may_be_replaced_by 0. +round_rules(25): word__last may_be_replaced_by 4294967295. +round_rules(26): word__base__first may_be_replaced_by 0. +round_rules(27): word__base__last may_be_replaced_by 4294967295. +round_rules(28): word__modulus may_be_replaced_by 4294967296. +round_rules(29): chain__size >= 0 may_be_deduced. +round_rules(30): A = B may_be_deduced_from + [goal(checktype(A,chain)), + goal(checktype(B,chain)), + fld_h0(A) = fld_h0(B), + fld_h1(A) = fld_h1(B), + fld_h2(A) = fld_h2(B), + fld_h3(A) = fld_h3(B), + fld_h4(A) = fld_h4(B)]. +round_rules(31): block_index__size >= 0 may_be_deduced. +round_rules(32): block_index__first may_be_replaced_by 0. +round_rules(33): block_index__last may_be_replaced_by 15. +round_rules(34): block_index__base__first <= block_index__base__last may_be_deduced. +round_rules(35): block_index__base__first <= block_index__first may_be_deduced. +round_rules(36): block_index__base__last >= block_index__last may_be_deduced. +round_rules(37): round_index__size >= 0 may_be_deduced. +round_rules(38): round_index__first may_be_replaced_by 0. +round_rules(39): round_index__last may_be_replaced_by 79. +round_rules(40): round_index__base__first <= round_index__base__last may_be_deduced. +round_rules(41): round_index__base__first <= round_index__first may_be_deduced. +round_rules(42): round_index__base__last >= round_index__last may_be_deduced. +round_rules(43): chain_pair__size >= 0 may_be_deduced. +round_rules(44): A = B may_be_deduced_from + [goal(checktype(A,chain_pair)), + goal(checktype(B,chain_pair)), + fld_left(A) = fld_left(B), + fld_right(A) = fld_right(B)]. +round_rules(45): rotate_amount__size >= 0 may_be_deduced. +round_rules(46): rotate_amount__first may_be_replaced_by 0. +round_rules(47): rotate_amount__last may_be_replaced_by 15. +round_rules(48): rotate_amount__base__first may_be_replaced_by -2147483648. +round_rules(49): rotate_amount__base__last may_be_replaced_by 2147483647. diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/round.siv --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/round.siv Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,834 @@ +***************************************************************************** + Semantic Analysis of SPARK Text + Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 + Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. +***************************************************************************** + + +CREATED 29-NOV-2010, 14:30:19 SIMPLIFIED 29-NOV-2010, 14:30:21 + +SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 +Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. + +procedure RMD.Round + + + + +For path(s) from start to run-time check associated with statement of line 111: + +procedure_round_1. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 112: + +procedure_round_2. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 113: + +procedure_round_3. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 114: + +procedure_round_4. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 115: + +procedure_round_5. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 116: + +procedure_round_6. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 117: + +procedure_round_7. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 118: + +procedure_round_8. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 119: + +procedure_round_9. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 120: + +procedure_round_10. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 121: + +procedure_round_11. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 121: + +procedure_round_12. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 124: + +procedure_round_13. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 124: + +procedure_round_14. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 124: + +procedure_round_15. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 124: + +procedure_round_16. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 124: + +procedure_round_17. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 124: + +procedure_round_18. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 124: + +procedure_round_19. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 124: + +procedure_round_20. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 124: + +procedure_round_21. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 124: + +procedure_round_22. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 124: + +procedure_round_23. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 124: + +procedure_round_24. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 130: + +procedure_round_25. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 130: + +procedure_round_26. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 131: + +procedure_round_27. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 131: + +procedure_round_28. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 132: + +procedure_round_29. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 132: + +procedure_round_30. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 132: + +procedure_round_31. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 132: + +procedure_round_32. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 133: + +procedure_round_33. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 133: + +procedure_round_34. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 134: + +procedure_round_35. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 134: + +procedure_round_36. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 136: + +procedure_round_37. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 136: + +procedure_round_38. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 136: + +procedure_round_39. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 136: + +procedure_round_40. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 136: + +procedure_round_41. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 136: + +procedure_round_42. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 136: + +procedure_round_43. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 136: + +procedure_round_44. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 136: + +procedure_round_45. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 136: + +procedure_round_46. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 136: + +procedure_round_47. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 136: + +procedure_round_48. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 142: + +procedure_round_49. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 142: + +procedure_round_50. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 143: + +procedure_round_51. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 143: + +procedure_round_52. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 144: + +procedure_round_53. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 144: + +procedure_round_54. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 144: + +procedure_round_55. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 144: + +procedure_round_56. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 145: + +procedure_round_57. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 145: + +procedure_round_58. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 146: + +procedure_round_59. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 146: + +procedure_round_60. +*** true . /* all conclusions proved */ + + +For path(s) from start to assertion of line 147: + +procedure_round_61. +H1: ca >= 0 . +H2: ca <= 4294967295 . +H3: cb >= 0 . +H4: cb <= 4294967295 . +H5: cc >= 0 . +H6: cc <= 4294967295 . +H7: cd >= 0 . +H8: cd <= 4294967295 . +H9: ce >= 0 . +H10: ce <= 4294967295 . +H11: for_all(i___1 : integer, 0 <= i___1 and i___1 <= 15 -> 0 <= element(x, [ + i___1]) and element(x, [i___1]) <= 4294967295) . +H12: s_l(0) >= 0 . +H13: s_l(0) <= 15 . +H14: s_l(0) = s_l_spec(0) . +H15: f(0, cb, cc, cd) >= 0 . +H16: f(0, cb, cc, cd) <= 4294967295 . +H17: f(0, cb, cc, cd) = f_spec(0, cb, cc, cd) . +H18: r_l(0) >= 0 . +H19: r_l(0) <= 15 . +H20: r_l(0) = r_l_spec(0) . +H21: k_l(0) >= 0 . +H22: k_l(0) <= 4294967295 . +H23: k_l(0) = k_l_spec(0) . +H24: (((ca + f(0, cb, cc, cd)) mod 4294967296 + element(x, [r_l(0)])) mod + 4294967296 + k_l(0)) mod 4294967296 >= 0 . +H25: (((ca + f(0, cb, cc, cd)) mod 4294967296 + element(x, [r_l(0)])) mod + 4294967296 + k_l(0)) mod 4294967296 <= 4294967295 . +H26: wordops__rotate(s_l(0), (((ca + f(0, cb, cc, cd)) mod 4294967296 + + element(x, [r_l(0)])) mod 4294967296 + k_l(0)) mod 4294967296) >= 0 . +H27: wordops__rotate(s_l(0), (((ca + f(0, cb, cc, cd)) mod 4294967296 + + element(x, [r_l(0)])) mod 4294967296 + k_l(0)) mod 4294967296) <= + 4294967295 . +H28: wordops__rotate(s_l(0), (((ca + f(0, cb, cc, cd)) mod 4294967296 + + element(x, [r_l(0)])) mod 4294967296 + k_l(0)) mod 4294967296) = + wordops__rotate_left(s_l(0), (((ca + f(0, cb, cc, cd)) mod 4294967296 + + element(x, [r_l(0)])) mod 4294967296 + k_l(0)) mod 4294967296) . +H29: (wordops__rotate(s_l(0), (((ca + f(0, cb, cc, cd)) mod 4294967296 + + element(x, [r_l(0)])) mod 4294967296 + k_l(0)) mod 4294967296) + ce) + mod 4294967296 >= 0 . +H30: (wordops__rotate(s_l(0), (((ca + f(0, cb, cc, cd)) mod 4294967296 + + element(x, [r_l(0)])) mod 4294967296 + k_l(0)) mod 4294967296) + ce) + mod 4294967296 <= 4294967295 . +H31: wordops__rotate(10, cc) >= 0 . +H32: wordops__rotate(10, cc) <= 4294967295 . +H33: wordops__rotate(10, cc) = wordops__rotate_left(10, cc) . +H34: s_r(0) >= 0 . +H35: s_r(0) <= 15 . +H36: s_r(0) = s_r_spec(0) . +H37: 79 >= round_index__base__first . +H38: 79 <= round_index__base__last . +H39: f(79, cb, cc, cd) >= 0 . +H40: f(79, cb, cc, cd) <= 4294967295 . +H41: f(79, cb, cc, cd) = f_spec(79, cb, cc, cd) . +H42: r_r(0) >= 0 . +H43: r_r(0) <= 15 . +H44: r_r(0) = r_r_spec(0) . +H45: k_r(0) >= 0 . +H46: k_r(0) <= 4294967295 . +H47: k_r(0) = k_r_spec(0) . +H48: (((ca + f(79, cb, cc, cd)) mod 4294967296 + element(x, [r_r(0)])) mod + 4294967296 + k_r(0)) mod 4294967296 >= 0 . +H49: (((ca + f(79, cb, cc, cd)) mod 4294967296 + element(x, [r_r(0)])) mod + 4294967296 + k_r(0)) mod 4294967296 <= 4294967295 . +H50: wordops__rotate(s_r(0), (((ca + f(79, cb, cc, cd)) mod 4294967296 + + element(x, [r_r(0)])) mod 4294967296 + k_r(0)) mod 4294967296) >= 0 . +H51: wordops__rotate(s_r(0), (((ca + f(79, cb, cc, cd)) mod 4294967296 + + element(x, [r_r(0)])) mod 4294967296 + k_r(0)) mod 4294967296) <= + 4294967295 . +H52: wordops__rotate(s_r(0), (((ca + f(79, cb, cc, cd)) mod 4294967296 + + element(x, [r_r(0)])) mod 4294967296 + k_r(0)) mod 4294967296) = + wordops__rotate_left(s_r(0), (((ca + f(79, cb, cc, cd)) mod + 4294967296 + element(x, [r_r(0)])) mod 4294967296 + k_r(0)) mod + 4294967296) . +H53: (wordops__rotate(s_r(0), (((ca + f(79, cb, cc, cd)) mod 4294967296 + + element(x, [r_r(0)])) mod 4294967296 + k_r(0)) mod 4294967296) + ce) + mod 4294967296 >= 0 . +H54: (wordops__rotate(s_r(0), (((ca + f(79, cb, cc, cd)) mod 4294967296 + + element(x, [r_r(0)])) mod 4294967296 + k_r(0)) mod 4294967296) + ce) + mod 4294967296 <= 4294967295 . +H55: integer__size >= 0 . +H56: interfaces__unsigned_32__size >= 0 . +H57: wordops__word__size >= 0 . +H58: wordops__rotate_amount__size >= 0 . +H59: word__size >= 0 . +H60: chain__size >= 0 . +H61: block_index__size >= 0 . +H62: block_index__base__first <= block_index__base__last . +H63: round_index__size >= 0 . +H64: round_index__base__first <= round_index__base__last . +H65: chain_pair__size >= 0 . +H66: rotate_amount__size >= 0 . +H67: block_index__base__first <= 0 . +H68: block_index__base__last >= 15 . +H69: round_index__base__first <= 0 . +H70: round_index__base__last >= 79 . + -> +C1: mk__chain_pair(left := mk__chain(h0 := ce, h1 := (wordops__rotate(s_l(0) + , (((ca + f(0, cb, cc, cd)) mod 4294967296 + element(x, [r_l(0)])) + mod 4294967296 + k_l(0)) mod 4294967296) + ce) mod 4294967296, h2 := + cb, h3 := wordops__rotate(10, cc), h4 := cd), right := mk__chain(h0 + := ce, h1 := (wordops__rotate(s_r(0), (((ca + f(79, cb, cc, cd)) mod + 4294967296 + element(x, [r_r(0)])) mod 4294967296 + k_r(0)) mod + 4294967296) + ce) mod 4294967296, h2 := cb, h3 := wordops__rotate(10, + cc), h4 := cd)) = steps(mk__chain_pair(left := mk__chain(h0 := ca, h1 + := cb, h2 := cc, h3 := cd, h4 := ce), right := mk__chain(h0 := ca, h1 + := cb, h2 := cc, h3 := cd, h4 := ce)), 1, x) . + + +For path(s) from assertion of line 147 to assertion of line 147: + +procedure_round_62. +H1: mk__chain_pair(left := mk__chain(h0 := cla, h1 := clb, h2 := clc, h3 := + cld, h4 := cle), right := mk__chain(h0 := cra, h1 := crb, h2 := crc, + h3 := crd, h4 := cre)) = steps(mk__chain_pair(left := mk__chain(h0 := + ca~, h1 := cb~, h2 := cc~, h3 := cd~, h4 := ce~), right := mk__chain( + h0 := ca~, h1 := cb~, h2 := cc~, h3 := cd~, h4 := ce~)), loop__1__j + + 1, x) . +H2: ca~ >= 0 . +H3: ca~ <= 4294967295 . +H4: cb~ >= 0 . +H5: cb~ <= 4294967295 . +H6: cc~ >= 0 . +H7: cc~ <= 4294967295 . +H8: cd~ >= 0 . +H9: cd~ <= 4294967295 . +H10: ce~ >= 0 . +H11: ce~ <= 4294967295 . +H12: for_all(i___1 : integer, 0 <= i___1 and i___1 <= 15 -> 0 <= element(x, [ + i___1]) and element(x, [i___1]) <= 4294967295) . +H13: loop__1__j >= 0 . +H14: loop__1__j <= 78 . +H15: s_l(loop__1__j + 1) >= 0 . +H16: s_l(loop__1__j + 1) <= 15 . +H17: s_l(loop__1__j + 1) = s_l_spec(loop__1__j + 1) . +H18: cla >= 0 . +H19: cla <= 4294967295 . +H20: clb >= 0 . +H21: clb <= 4294967295 . +H22: clc >= 0 . +H23: clc <= 4294967295 . +H24: cld >= 0 . +H25: cld <= 4294967295 . +H26: f(loop__1__j + 1, clb, clc, cld) >= 0 . +H27: f(loop__1__j + 1, clb, clc, cld) <= 4294967295 . +H28: f(loop__1__j + 1, clb, clc, cld) = f_spec(loop__1__j + 1, clb, clc, cld) + . +H29: r_l(loop__1__j + 1) >= 0 . +H30: r_l(loop__1__j + 1) <= 15 . +H31: r_l(loop__1__j + 1) = r_l_spec(loop__1__j + 1) . +H32: k_l(loop__1__j + 1) >= 0 . +H33: k_l(loop__1__j + 1) <= 4294967295 . +H34: k_l(loop__1__j + 1) = k_l_spec(loop__1__j + 1) . +H35: (((cla + f(loop__1__j + 1, clb, clc, cld)) mod 4294967296 + element(x, [ + r_l(loop__1__j + 1)])) mod 4294967296 + k_l(loop__1__j + 1)) mod + 4294967296 >= 0 . +H36: (((cla + f(loop__1__j + 1, clb, clc, cld)) mod 4294967296 + element(x, [ + r_l(loop__1__j + 1)])) mod 4294967296 + k_l(loop__1__j + 1)) mod + 4294967296 <= 4294967295 . +H37: wordops__rotate(s_l(loop__1__j + 1), (((cla + f(loop__1__j + 1, clb, + clc, cld)) mod 4294967296 + element(x, [r_l(loop__1__j + 1)])) mod + 4294967296 + k_l(loop__1__j + 1)) mod 4294967296) >= 0 . +H38: wordops__rotate(s_l(loop__1__j + 1), (((cla + f(loop__1__j + 1, clb, + clc, cld)) mod 4294967296 + element(x, [r_l(loop__1__j + 1)])) mod + 4294967296 + k_l(loop__1__j + 1)) mod 4294967296) <= 4294967295 . +H39: wordops__rotate(s_l(loop__1__j + 1), (((cla + f(loop__1__j + 1, clb, + clc, cld)) mod 4294967296 + element(x, [r_l(loop__1__j + 1)])) mod + 4294967296 + k_l(loop__1__j + 1)) mod 4294967296) = + wordops__rotate_left(s_l(loop__1__j + 1), (((cla + f(loop__1__j + 1, + clb, clc, cld)) mod 4294967296 + element(x, [r_l(loop__1__j + 1)])) + mod 4294967296 + k_l(loop__1__j + 1)) mod 4294967296) . +H40: cle >= 0 . +H41: cle <= 4294967295 . +H42: (wordops__rotate(s_l(loop__1__j + 1), (((cla + f(loop__1__j + 1, clb, + clc, cld)) mod 4294967296 + element(x, [r_l(loop__1__j + 1)])) mod + 4294967296 + k_l(loop__1__j + 1)) mod 4294967296) + cle) mod + 4294967296 >= 0 . +H43: (wordops__rotate(s_l(loop__1__j + 1), (((cla + f(loop__1__j + 1, clb, + clc, cld)) mod 4294967296 + element(x, [r_l(loop__1__j + 1)])) mod + 4294967296 + k_l(loop__1__j + 1)) mod 4294967296) + cle) mod + 4294967296 <= 4294967295 . +H44: wordops__rotate(10, clc) >= 0 . +H45: wordops__rotate(10, clc) <= 4294967295 . +H46: wordops__rotate(10, clc) = wordops__rotate_left(10, clc) . +H47: s_r(loop__1__j + 1) >= 0 . +H48: s_r(loop__1__j + 1) <= 15 . +H49: s_r(loop__1__j + 1) = s_r_spec(loop__1__j + 1) . +H50: cra >= 0 . +H51: cra <= 4294967295 . +H52: crb >= 0 . +H53: crb <= 4294967295 . +H54: crc >= 0 . +H55: crc <= 4294967295 . +H56: crd >= 0 . +H57: crd <= 4294967295 . +H58: 79 - (loop__1__j + 1) >= round_index__base__first . +H59: 79 - (loop__1__j + 1) <= round_index__base__last . +H60: f(79 - (loop__1__j + 1), crb, crc, crd) >= 0 . +H61: f(79 - (loop__1__j + 1), crb, crc, crd) <= 4294967295 . +H62: f(78 - loop__1__j, crb, crc, crd) = f_spec(78 - loop__1__j, crb, crc, + crd) . +H63: r_r(loop__1__j + 1) >= 0 . +H64: r_r(loop__1__j + 1) <= 15 . +H65: r_r(loop__1__j + 1) = r_r_spec(loop__1__j + 1) . +H66: k_r(loop__1__j + 1) >= 0 . +H67: k_r(loop__1__j + 1) <= 4294967295 . +H68: k_r(loop__1__j + 1) = k_r_spec(loop__1__j + 1) . +H69: (((cra + f(79 - (loop__1__j + 1), crb, crc, crd)) mod 4294967296 + + element(x, [r_r(loop__1__j + 1)])) mod 4294967296 + k_r(loop__1__j + + 1)) mod 4294967296 >= 0 . +H70: (((cra + f(79 - (loop__1__j + 1), crb, crc, crd)) mod 4294967296 + + element(x, [r_r(loop__1__j + 1)])) mod 4294967296 + k_r(loop__1__j + + 1)) mod 4294967296 <= 4294967295 . +H71: wordops__rotate(s_r(loop__1__j + 1), (((cra + f(79 - (loop__1__j + 1), + crb, crc, crd)) mod 4294967296 + element(x, [r_r(loop__1__j + 1)])) + mod 4294967296 + k_r(loop__1__j + 1)) mod 4294967296) >= 0 . +H72: wordops__rotate(s_r(loop__1__j + 1), (((cra + f(79 - (loop__1__j + 1), + crb, crc, crd)) mod 4294967296 + element(x, [r_r(loop__1__j + 1)])) + mod 4294967296 + k_r(loop__1__j + 1)) mod 4294967296) <= 4294967295 . +H73: wordops__rotate(s_r(loop__1__j + 1), (((cra + f(79 - (loop__1__j + 1), + crb, crc, crd)) mod 4294967296 + element(x, [r_r(loop__1__j + 1)])) + mod 4294967296 + k_r(loop__1__j + 1)) mod 4294967296) = + wordops__rotate_left(s_r(loop__1__j + 1), (((cra + f(79 - (loop__1__j + + 1), crb, crc, crd)) mod 4294967296 + element(x, [r_r(loop__1__j + 1) + ])) mod 4294967296 + k_r(loop__1__j + 1)) mod 4294967296) . +H74: cre >= 0 . +H75: cre <= 4294967295 . +H76: (wordops__rotate(s_r(loop__1__j + 1), (((cra + f(79 - (loop__1__j + 1), + crb, crc, crd)) mod 4294967296 + element(x, [r_r(loop__1__j + 1)])) + mod 4294967296 + k_r(loop__1__j + 1)) mod 4294967296) + cre) mod + 4294967296 >= 0 . +H77: (wordops__rotate(s_r(loop__1__j + 1), (((cra + f(79 - (loop__1__j + 1), + crb, crc, crd)) mod 4294967296 + element(x, [r_r(loop__1__j + 1)])) + mod 4294967296 + k_r(loop__1__j + 1)) mod 4294967296) + cre) mod + 4294967296 <= 4294967295 . +H78: wordops__rotate(10, crc) >= 0 . +H79: wordops__rotate(10, crc) <= 4294967295 . +H80: wordops__rotate(10, crc) = wordops__rotate_left(10, crc) . +H81: integer__size >= 0 . +H82: interfaces__unsigned_32__size >= 0 . +H83: wordops__word__size >= 0 . +H84: wordops__rotate_amount__size >= 0 . +H85: word__size >= 0 . +H86: chain__size >= 0 . +H87: block_index__size >= 0 . +H88: block_index__base__first <= block_index__base__last . +H89: round_index__size >= 0 . +H90: round_index__base__first <= round_index__base__last . +H91: chain_pair__size >= 0 . +H92: rotate_amount__size >= 0 . +H93: block_index__base__first <= 0 . +H94: block_index__base__last >= 15 . +H95: round_index__base__first <= 0 . +H96: round_index__base__last >= 79 . + -> +C1: mk__chain_pair(left := mk__chain(h0 := cle, h1 := (wordops__rotate(s_l( + loop__1__j + 1), (((cla + f(loop__1__j + 1, clb, clc, cld)) mod + 4294967296 + element(x, [r_l(loop__1__j + 1)])) mod 4294967296 + k_l( + loop__1__j + 1)) mod 4294967296) + cle) mod 4294967296, h2 := clb, h3 + := wordops__rotate(10, clc), h4 := cld), right := mk__chain(h0 := + cre, h1 := (wordops__rotate(s_r(loop__1__j + 1), (((cra + f(79 - ( + loop__1__j + 1), crb, crc, crd)) mod 4294967296 + element(x, [r_r( + loop__1__j + 1)])) mod 4294967296 + k_r(loop__1__j + 1)) mod + 4294967296) + cre) mod 4294967296, h2 := crb, h3 := wordops__rotate( + 10, crc), h4 := crd)) = steps(mk__chain_pair(left := mk__chain(h0 := + ca~, h1 := cb~, h2 := cc~, h3 := cd~, h4 := ce~), right := mk__chain( + h0 := ca~, h1 := cb~, h2 := cc~, h3 := cd~, h4 := ce~)), loop__1__j + + 2, x) . + + +For path(s) from start to run-time check associated with statement of line 153: + +procedure_round_63. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 153: + +procedure_round_64. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 154: + +procedure_round_65. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 154: + +procedure_round_66. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 155: + +procedure_round_67. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 155: + +procedure_round_68. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 156: + +procedure_round_69. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 156: + +procedure_round_70. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 157: + +procedure_round_71. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 157: + +procedure_round_72. +*** true . /* all conclusions proved */ + + +For path(s) from start to run-time check associated with statement of line 158: + +procedure_round_73. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 147 to run-time check associated with + statement of line 158: + +procedure_round_74. +*** true . /* all conclusions proved */ + + +For path(s) from start to finish: + +procedure_round_75. +*** true . /* contradiction within hypotheses. */ + + + +For path(s) from assertion of line 147 to finish: + +procedure_round_76. +H1: mk__chain_pair(left := mk__chain(h0 := cla, h1 := clb, h2 := clc, h3 := + cld, h4 := cle), right := mk__chain(h0 := cra, h1 := crb, h2 := crc, + h3 := crd, h4 := cre)) = steps(mk__chain_pair(left := mk__chain(h0 := + ca~, h1 := cb~, h2 := cc~, h3 := cd~, h4 := ce~), right := mk__chain( + h0 := ca~, h1 := cb~, h2 := cc~, h3 := cd~, h4 := ce~)), 80, x) . +H2: ca~ >= 0 . +H3: ca~ <= 4294967295 . +H4: cb~ >= 0 . +H5: cb~ <= 4294967295 . +H6: cc~ >= 0 . +H7: cc~ <= 4294967295 . +H8: cd~ >= 0 . +H9: cd~ <= 4294967295 . +H10: ce~ >= 0 . +H11: ce~ <= 4294967295 . +H12: for_all(i___1 : integer, 0 <= i___1 and i___1 <= 15 -> 0 <= element(x, [ + i___1]) and element(x, [i___1]) <= 4294967295) . +H13: clc >= 0 . +H14: clc <= 4294967295 . +H15: crd >= 0 . +H16: crd <= 4294967295 . +H17: ((cb~ + clc) mod 4294967296 + crd) mod 4294967296 >= 0 . +H18: ((cb~ + clc) mod 4294967296 + crd) mod 4294967296 <= 4294967295 . +H19: cld >= 0 . +H20: cld <= 4294967295 . +H21: cre >= 0 . +H22: cre <= 4294967295 . +H23: ((cc~ + cld) mod 4294967296 + cre) mod 4294967296 >= 0 . +H24: ((cc~ + cld) mod 4294967296 + cre) mod 4294967296 <= 4294967295 . +H25: cle >= 0 . +H26: cle <= 4294967295 . +H27: cra >= 0 . +H28: cra <= 4294967295 . +H29: ((cd~ + cle) mod 4294967296 + cra) mod 4294967296 >= 0 . +H30: ((cd~ + cle) mod 4294967296 + cra) mod 4294967296 <= 4294967295 . +H31: cla >= 0 . +H32: cla <= 4294967295 . +H33: crb >= 0 . +H34: crb <= 4294967295 . +H35: ((ce~ + cla) mod 4294967296 + crb) mod 4294967296 >= 0 . +H36: ((ce~ + cla) mod 4294967296 + crb) mod 4294967296 <= 4294967295 . +H37: clb >= 0 . +H38: clb <= 4294967295 . +H39: crc >= 0 . +H40: crc <= 4294967295 . +H41: ((ca~ + clb) mod 4294967296 + crc) mod 4294967296 >= 0 . +H42: ((ca~ + clb) mod 4294967296 + crc) mod 4294967296 <= 4294967295 . +H43: integer__size >= 0 . +H44: interfaces__unsigned_32__size >= 0 . +H45: wordops__word__size >= 0 . +H46: wordops__rotate_amount__size >= 0 . +H47: word__size >= 0 . +H48: chain__size >= 0 . +H49: block_index__size >= 0 . +H50: block_index__base__first <= block_index__base__last . +H51: round_index__size >= 0 . +H52: round_index__base__first <= round_index__base__last . +H53: chain_pair__size >= 0 . +H54: rotate_amount__size >= 0 . +H55: block_index__base__first <= 0 . +H56: block_index__base__last >= 15 . +H57: round_index__base__first <= 0 . +H58: round_index__base__last >= 79 . + -> +C1: mk__chain(h0 := ((cb~ + clc) mod 4294967296 + crd) mod 4294967296, h1 := + ((cc~ + cld) mod 4294967296 + cre) mod 4294967296, h2 := ((cd~ + cle) + mod 4294967296 + cra) mod 4294967296, h3 := ((ce~ + cla) mod + 4294967296 + crb) mod 4294967296, h4 := ((ca~ + clb) mod 4294967296 + + crc) mod 4294967296) = round_spec(mk__chain(h0 := ca~, h1 := cb~, h2 + := cc~, h3 := cd~, h4 := ce~), x) . + + diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_l.fdl --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_l.fdl Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,37 @@ + {*******************************************************} + {FDL Declarations} + {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039} + {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.} + {*******************************************************} + + + {DATE : 29-NOV-2010 14:30:19.83} + + {function RMD.S_L} + + +title function s_l; + + function round__(real) : integer; + type round_index = integer; + type rotate_definition = array [integer] of integer; + const s_values : rotate_definition = pending; + const rotate_amount__base__first : integer = pending; + const rotate_amount__base__last : integer = pending; + const round_index__base__first : integer = pending; + const round_index__base__last : integer = pending; + const integer__base__first : integer = pending; + const integer__base__last : integer = pending; + const rotate_amount__first : integer = pending; + const rotate_amount__last : integer = pending; + const rotate_amount__size : integer = pending; + const round_index__first : integer = pending; + const round_index__last : integer = pending; + const round_index__size : integer = pending; + const integer__first : integer = pending; + const integer__last : integer = pending; + const integer__size : integer = pending; + var j : integer; + function s_l_spec(integer) : integer; + +end; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_l.rls --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_l.rls Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,81 @@ + /*********************************************************/ + /*Proof Rule Declarations*/ + /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/ + /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/ + /*********************************************************/ + + + /*DATE : 29-NOV-2010 14:30:19.83*/ + + /*function RMD.S_L*/ + + +rule_family s_l_rules: + X requires [X:any] & + X <= Y requires [X:ire, Y:ire] & + X >= Y requires [X:ire, Y:ire]. + +s_l_rules(1): rotate_amount__first <= element(s_values, [I]) may_be_deduced_from [0 <= I, I <= 79]. +s_l_rules(2): element(s_values, [I]) <= rotate_amount__last may_be_deduced_from [0 <= I, I <= 79]. +s_l_rules(3): s_values may_be_replaced_by + mk__rotate_definition([round_index__first] := 11, [ + round_index__first + 1] := 14, [round_index__first + 2] := + 15, [round_index__first + 3] := 12, [round_index__first + 4] := + 5, [round_index__first + 5] := 8, [round_index__first + 6] := + 7, [round_index__first + 7] := 9, [round_index__first + 8] := + 11, [round_index__first + 9] := 13, [round_index__first + 10] := + 14, [round_index__first + 11] := 15, [ + round_index__first + 12] := 6, [round_index__first + 13] := + 7, [round_index__first + 14] := 9, [round_index__first + 15] := + 8, [round_index__first + 16] := 7, [round_index__first + 17] := + 6, [round_index__first + 18] := 8, [round_index__first + 19] := + 13, [round_index__first + 20] := 11, [ + round_index__first + 21] := 9, [round_index__first + 22] := + 7, [round_index__first + 23] := 15, [round_index__first + 24] := + 7, [round_index__first + 25] := 12, [round_index__first + 26] := + 15, [round_index__first + 27] := 9, [round_index__first + 28] := + 11, [round_index__first + 29] := 7, [round_index__first + 30] := + 13, [round_index__first + 31] := 12, [ + round_index__first + 32] := 11, [round_index__first + 33] := + 13, [round_index__first + 34] := 6, [round_index__first + 35] := + 7, [round_index__first + 36] := 14, [round_index__first + 37] := + 9, [round_index__first + 38] := 13, [round_index__first + 39] := + 15, [round_index__first + 40] := 14, [ + round_index__first + 41] := 8, [round_index__first + 42] := + 13, [round_index__first + 43] := 6, [round_index__first + 44] := + 5, [round_index__first + 45] := 12, [round_index__first + 46] := + 7, [round_index__first + 47] := 5, [round_index__first + 48] := + 11, [round_index__first + 49] := 12, [ + round_index__first + 50] := 14, [round_index__first + 51] := + 15, [round_index__first + 52] := 14, [ + round_index__first + 53] := 15, [round_index__first + 54] := + 9, [round_index__first + 55] := 8, [round_index__first + 56] := + 9, [round_index__first + 57] := 14, [round_index__first + 58] := + 5, [round_index__first + 59] := 6, [round_index__first + 60] := + 8, [round_index__first + 61] := 6, [round_index__first + 62] := + 5, [round_index__first + 63] := 12, [round_index__first + 64] := + 9, [round_index__first + 65] := 15, [round_index__first + 66] := + 5, [round_index__first + 67] := 11, [round_index__first + 68] := + 6, [round_index__first + 69] := 8, [round_index__first + 70] := + 13, [round_index__first + 71] := 12, [ + round_index__first + 72] := 5, [round_index__first + 73] := + 12, [round_index__first + 74] := 13, [ + round_index__first + 75] := 14, [round_index__first + 76] := + 11, [round_index__first + 77] := 8, [round_index__first + 78] := + 5, [round_index__first + 79] := 6). +s_l_rules(4): integer__size >= 0 may_be_deduced. +s_l_rules(5): integer__first may_be_replaced_by -2147483648. +s_l_rules(6): integer__last may_be_replaced_by 2147483647. +s_l_rules(7): integer__base__first may_be_replaced_by -2147483648. +s_l_rules(8): integer__base__last may_be_replaced_by 2147483647. +s_l_rules(9): round_index__size >= 0 may_be_deduced. +s_l_rules(10): round_index__first may_be_replaced_by 0. +s_l_rules(11): round_index__last may_be_replaced_by 79. +s_l_rules(12): round_index__base__first <= round_index__base__last may_be_deduced. +s_l_rules(13): round_index__base__first <= round_index__first may_be_deduced. +s_l_rules(14): round_index__base__last >= round_index__last may_be_deduced. +s_l_rules(15): rotate_amount__size >= 0 may_be_deduced. +s_l_rules(16): rotate_amount__first may_be_replaced_by 0. +s_l_rules(17): rotate_amount__last may_be_replaced_by 15. +s_l_rules(18): rotate_amount__base__first may_be_replaced_by -2147483648. +s_l_rules(19): rotate_amount__base__last may_be_replaced_by 2147483647. diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_l.siv --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_l.siv Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,79 @@ +***************************************************************************** + Semantic Analysis of SPARK Text + Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 + Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. +***************************************************************************** + + +CREATED 29-NOV-2010, 14:30:19 SIMPLIFIED 29-NOV-2010, 14:30:29 + +SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 +Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. + +function RMD.S_L + + + + +For path(s) from start to run-time check associated with statement of line 87: + +function_s_l_1. +*** true . /* all conclusions proved */ + + +For path(s) from start to finish: + +function_s_l_2. +H1: j >= 0 . +H2: j <= 79 . +H3: integer__size >= 0 . +H4: round_index__size >= 0 . +H5: round_index__base__first <= round_index__base__last . +H6: rotate_amount__size >= 0 . +H7: round_index__base__first <= 0 . +H8: round_index__base__last >= 79 . + -> +C1: element(mk__rotate_definition([0] := 11, [1] := 14, [2] := 15, [3] := + 12, [4] := 5, [5] := 8, [6] := 7, [7] := 9, [8] := 11, [9] := 13, [10] + := 14, [11] := 15, [12] := 6, [13] := 7, [14] := 9, [15] := 8, [16] + := 7, [17] := 6, [18] := 8, [19] := 13, [20] := 11, [21] := 9, [22] + := 7, [23] := 15, [24] := 7, [25] := 12, [26] := 15, [27] := 9, [28] + := 11, [29] := 7, [30] := 13, [31] := 12, [32] := 11, [33] := 13, [34] + := 6, [35] := 7, [36] := 14, [37] := 9, [38] := 13, [39] := 15, [40] + := 14, [41] := 8, [42] := 13, [43] := 6, [44] := 5, [45] := 12, [46] + := 7, [47] := 5, [48] := 11, [49] := 12, [50] := 14, [51] := 15, [52] + := 14, [53] := 15, [54] := 9, [55] := 8, [56] := 9, [57] := 14, [58] + := 5, [59] := 6, [60] := 8, [61] := 6, [62] := 5, [63] := 12, [64] := + 9, [65] := 15, [66] := 5, [67] := 11, [68] := 6, [69] := 8, [70] := + 13, [71] := 12, [72] := 5, [73] := 12, [74] := 13, [75] := 14, [76] + := 11, [77] := 8, [78] := 5, [79] := 6), [j]) = s_l_spec(j) . +C2: element(mk__rotate_definition([0] := 11, [1] := 14, [2] := 15, [3] := + 12, [4] := 5, [5] := 8, [6] := 7, [7] := 9, [8] := 11, [9] := 13, [10] + := 14, [11] := 15, [12] := 6, [13] := 7, [14] := 9, [15] := 8, [16] + := 7, [17] := 6, [18] := 8, [19] := 13, [20] := 11, [21] := 9, [22] + := 7, [23] := 15, [24] := 7, [25] := 12, [26] := 15, [27] := 9, [28] + := 11, [29] := 7, [30] := 13, [31] := 12, [32] := 11, [33] := 13, [34] + := 6, [35] := 7, [36] := 14, [37] := 9, [38] := 13, [39] := 15, [40] + := 14, [41] := 8, [42] := 13, [43] := 6, [44] := 5, [45] := 12, [46] + := 7, [47] := 5, [48] := 11, [49] := 12, [50] := 14, [51] := 15, [52] + := 14, [53] := 15, [54] := 9, [55] := 8, [56] := 9, [57] := 14, [58] + := 5, [59] := 6, [60] := 8, [61] := 6, [62] := 5, [63] := 12, [64] := + 9, [65] := 15, [66] := 5, [67] := 11, [68] := 6, [69] := 8, [70] := + 13, [71] := 12, [72] := 5, [73] := 12, [74] := 13, [75] := 14, [76] + := 11, [77] := 8, [78] := 5, [79] := 6), [j]) >= 0 . +C3: element(mk__rotate_definition([0] := 11, [1] := 14, [2] := 15, [3] := + 12, [4] := 5, [5] := 8, [6] := 7, [7] := 9, [8] := 11, [9] := 13, [10] + := 14, [11] := 15, [12] := 6, [13] := 7, [14] := 9, [15] := 8, [16] + := 7, [17] := 6, [18] := 8, [19] := 13, [20] := 11, [21] := 9, [22] + := 7, [23] := 15, [24] := 7, [25] := 12, [26] := 15, [27] := 9, [28] + := 11, [29] := 7, [30] := 13, [31] := 12, [32] := 11, [33] := 13, [34] + := 6, [35] := 7, [36] := 14, [37] := 9, [38] := 13, [39] := 15, [40] + := 14, [41] := 8, [42] := 13, [43] := 6, [44] := 5, [45] := 12, [46] + := 7, [47] := 5, [48] := 11, [49] := 12, [50] := 14, [51] := 15, [52] + := 14, [53] := 15, [54] := 9, [55] := 8, [56] := 9, [57] := 14, [58] + := 5, [59] := 6, [60] := 8, [61] := 6, [62] := 5, [63] := 12, [64] := + 9, [65] := 15, [66] := 5, [67] := 11, [68] := 6, [69] := 8, [70] := + 13, [71] := 12, [72] := 5, [73] := 12, [74] := 13, [75] := 14, [76] + := 11, [77] := 8, [78] := 5, [79] := 6), [j]) <= 15 . + + diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_r.fdl --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_r.fdl Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,37 @@ + {*******************************************************} + {FDL Declarations} + {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039} + {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.} + {*******************************************************} + + + {DATE : 29-NOV-2010 14:30:19.84} + + {function RMD.S_R} + + +title function s_r; + + function round__(real) : integer; + type round_index = integer; + type rotate_definition = array [integer] of integer; + const s_values : rotate_definition = pending; + const rotate_amount__base__first : integer = pending; + const rotate_amount__base__last : integer = pending; + const round_index__base__first : integer = pending; + const round_index__base__last : integer = pending; + const integer__base__first : integer = pending; + const integer__base__last : integer = pending; + const rotate_amount__first : integer = pending; + const rotate_amount__last : integer = pending; + const rotate_amount__size : integer = pending; + const round_index__first : integer = pending; + const round_index__last : integer = pending; + const round_index__size : integer = pending; + const integer__first : integer = pending; + const integer__last : integer = pending; + const integer__size : integer = pending; + var j : integer; + function s_r_spec(integer) : integer; + +end; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_r.rls --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_r.rls Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,81 @@ + /*********************************************************/ + /*Proof Rule Declarations*/ + /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/ + /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/ + /*********************************************************/ + + + /*DATE : 29-NOV-2010 14:30:19.84*/ + + /*function RMD.S_R*/ + + +rule_family s_r_rules: + X requires [X:any] & + X <= Y requires [X:ire, Y:ire] & + X >= Y requires [X:ire, Y:ire]. + +s_r_rules(1): rotate_amount__first <= element(s_values, [I]) may_be_deduced_from [0 <= I, I <= 79]. +s_r_rules(2): element(s_values, [I]) <= rotate_amount__last may_be_deduced_from [0 <= I, I <= 79]. +s_r_rules(3): s_values may_be_replaced_by + mk__rotate_definition([round_index__first] := 8, [ + round_index__first + 1] := 9, [round_index__first + 2] := 9, [ + round_index__first + 3] := 11, [round_index__first + 4] := + 13, [round_index__first + 5] := 15, [round_index__first + 6] := + 15, [round_index__first + 7] := 5, [round_index__first + 8] := + 7, [round_index__first + 9] := 7, [round_index__first + 10] := + 8, [round_index__first + 11] := 11, [round_index__first + 12] := + 14, [round_index__first + 13] := 14, [ + round_index__first + 14] := 12, [round_index__first + 15] := + 6, [round_index__first + 16] := 9, [round_index__first + 17] := + 13, [round_index__first + 18] := 15, [ + round_index__first + 19] := 7, [round_index__first + 20] := + 12, [round_index__first + 21] := 8, [round_index__first + 22] := + 9, [round_index__first + 23] := 11, [round_index__first + 24] := + 7, [round_index__first + 25] := 7, [round_index__first + 26] := + 12, [round_index__first + 27] := 7, [round_index__first + 28] := + 6, [round_index__first + 29] := 15, [round_index__first + 30] := + 13, [round_index__first + 31] := 11, [ + round_index__first + 32] := 9, [round_index__first + 33] := + 7, [round_index__first + 34] := 15, [round_index__first + 35] := + 11, [round_index__first + 36] := 8, [round_index__first + 37] := + 6, [round_index__first + 38] := 6, [round_index__first + 39] := + 14, [round_index__first + 40] := 12, [ + round_index__first + 41] := 13, [round_index__first + 42] := + 5, [round_index__first + 43] := 14, [round_index__first + 44] := + 13, [round_index__first + 45] := 13, [ + round_index__first + 46] := 7, [round_index__first + 47] := + 5, [round_index__first + 48] := 15, [round_index__first + 49] := + 5, [round_index__first + 50] := 8, [round_index__first + 51] := + 11, [round_index__first + 52] := 14, [ + round_index__first + 53] := 14, [round_index__first + 54] := + 6, [round_index__first + 55] := 14, [round_index__first + 56] := + 6, [round_index__first + 57] := 9, [round_index__first + 58] := + 12, [round_index__first + 59] := 9, [round_index__first + 60] := + 12, [round_index__first + 61] := 5, [round_index__first + 62] := + 15, [round_index__first + 63] := 8, [round_index__first + 64] := + 8, [round_index__first + 65] := 5, [round_index__first + 66] := + 12, [round_index__first + 67] := 9, [round_index__first + 68] := + 12, [round_index__first + 69] := 5, [round_index__first + 70] := + 14, [round_index__first + 71] := 6, [round_index__first + 72] := + 8, [round_index__first + 73] := 13, [round_index__first + 74] := + 6, [round_index__first + 75] := 5, [round_index__first + 76] := + 15, [round_index__first + 77] := 13, [ + round_index__first + 78] := 11, [round_index__first + 79] := + 11). +s_r_rules(4): integer__size >= 0 may_be_deduced. +s_r_rules(5): integer__first may_be_replaced_by -2147483648. +s_r_rules(6): integer__last may_be_replaced_by 2147483647. +s_r_rules(7): integer__base__first may_be_replaced_by -2147483648. +s_r_rules(8): integer__base__last may_be_replaced_by 2147483647. +s_r_rules(9): round_index__size >= 0 may_be_deduced. +s_r_rules(10): round_index__first may_be_replaced_by 0. +s_r_rules(11): round_index__last may_be_replaced_by 79. +s_r_rules(12): round_index__base__first <= round_index__base__last may_be_deduced. +s_r_rules(13): round_index__base__first <= round_index__first may_be_deduced. +s_r_rules(14): round_index__base__last >= round_index__last may_be_deduced. +s_r_rules(15): rotate_amount__size >= 0 may_be_deduced. +s_r_rules(16): rotate_amount__first may_be_replaced_by 0. +s_r_rules(17): rotate_amount__last may_be_replaced_by 15. +s_r_rules(18): rotate_amount__base__first may_be_replaced_by -2147483648. +s_r_rules(19): rotate_amount__base__last may_be_replaced_by 2147483647. diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_r.siv --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_r.siv Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,79 @@ +***************************************************************************** + Semantic Analysis of SPARK Text + Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 + Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. +***************************************************************************** + + +CREATED 29-NOV-2010, 14:30:19 SIMPLIFIED 29-NOV-2010, 14:30:30 + +SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 +Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. + +function RMD.S_R + + + + +For path(s) from start to run-time check associated with statement of line 101: + +function_s_r_1. +*** true . /* all conclusions proved */ + + +For path(s) from start to finish: + +function_s_r_2. +H1: j >= 0 . +H2: j <= 79 . +H3: integer__size >= 0 . +H4: round_index__size >= 0 . +H5: round_index__base__first <= round_index__base__last . +H6: rotate_amount__size >= 0 . +H7: round_index__base__first <= 0 . +H8: round_index__base__last >= 79 . + -> +C1: element(mk__rotate_definition([0] := 8, [1] := 9, [2] := 9, [3] := 11, [ + 4] := 13, [5] := 15, [6] := 15, [7] := 5, [8] := 7, [9] := 7, [10] := + 8, [11] := 11, [12] := 14, [13] := 14, [14] := 12, [15] := 6, [16] := + 9, [17] := 13, [18] := 15, [19] := 7, [20] := 12, [21] := 8, [22] := + 9, [23] := 11, [24] := 7, [25] := 7, [26] := 12, [27] := 7, [28] := + 6, [29] := 15, [30] := 13, [31] := 11, [32] := 9, [33] := 7, [34] := + 15, [35] := 11, [36] := 8, [37] := 6, [38] := 6, [39] := 14, [40] := + 12, [41] := 13, [42] := 5, [43] := 14, [44] := 13, [45] := 13, [46] + := 7, [47] := 5, [48] := 15, [49] := 5, [50] := 8, [51] := 11, [52] + := 14, [53] := 14, [54] := 6, [55] := 14, [56] := 6, [57] := 9, [58] + := 12, [59] := 9, [60] := 12, [61] := 5, [62] := 15, [63] := 8, [64] + := 8, [65] := 5, [66] := 12, [67] := 9, [68] := 12, [69] := 5, [70] + := 14, [71] := 6, [72] := 8, [73] := 13, [74] := 6, [75] := 5, [76] + := 15, [77] := 13, [78] := 11, [79] := 11), [j]) = s_r_spec(j) . +C2: element(mk__rotate_definition([0] := 8, [1] := 9, [2] := 9, [3] := 11, [ + 4] := 13, [5] := 15, [6] := 15, [7] := 5, [8] := 7, [9] := 7, [10] := + 8, [11] := 11, [12] := 14, [13] := 14, [14] := 12, [15] := 6, [16] := + 9, [17] := 13, [18] := 15, [19] := 7, [20] := 12, [21] := 8, [22] := + 9, [23] := 11, [24] := 7, [25] := 7, [26] := 12, [27] := 7, [28] := + 6, [29] := 15, [30] := 13, [31] := 11, [32] := 9, [33] := 7, [34] := + 15, [35] := 11, [36] := 8, [37] := 6, [38] := 6, [39] := 14, [40] := + 12, [41] := 13, [42] := 5, [43] := 14, [44] := 13, [45] := 13, [46] + := 7, [47] := 5, [48] := 15, [49] := 5, [50] := 8, [51] := 11, [52] + := 14, [53] := 14, [54] := 6, [55] := 14, [56] := 6, [57] := 9, [58] + := 12, [59] := 9, [60] := 12, [61] := 5, [62] := 15, [63] := 8, [64] + := 8, [65] := 5, [66] := 12, [67] := 9, [68] := 12, [69] := 5, [70] + := 14, [71] := 6, [72] := 8, [73] := 13, [74] := 6, [75] := 5, [76] + := 15, [77] := 13, [78] := 11, [79] := 11), [j]) >= 0 . +C3: element(mk__rotate_definition([0] := 8, [1] := 9, [2] := 9, [3] := 11, [ + 4] := 13, [5] := 15, [6] := 15, [7] := 5, [8] := 7, [9] := 7, [10] := + 8, [11] := 11, [12] := 14, [13] := 14, [14] := 12, [15] := 6, [16] := + 9, [17] := 13, [18] := 15, [19] := 7, [20] := 12, [21] := 8, [22] := + 9, [23] := 11, [24] := 7, [25] := 7, [26] := 12, [27] := 7, [28] := + 6, [29] := 15, [30] := 13, [31] := 11, [32] := 9, [33] := 7, [34] := + 15, [35] := 11, [36] := 8, [37] := 6, [38] := 6, [39] := 14, [40] := + 12, [41] := 13, [42] := 5, [43] := 14, [44] := 13, [45] := 13, [46] + := 7, [47] := 5, [48] := 15, [49] := 5, [50] := 8, [51] := 11, [52] + := 14, [53] := 14, [54] := 6, [55] := 14, [56] := 6, [57] := 9, [58] + := 12, [59] := 9, [60] := 12, [61] := 5, [62] := 15, [63] := 8, [64] + := 8, [65] := 5, [66] := 12, [67] := 9, [68] := 12, [69] := 5, [70] + := 14, [71] := 6, [72] := 8, [73] := 13, [74] := 6, [75] := 5, [76] + := 15, [77] := 13, [78] := 11, [79] := 11), [j]) <= 15 . + + diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/shadow/interfaces.ads --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/shadow/interfaces.ads Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,5 @@ +package Interfaces is + + type Unsigned_32 is mod 2 ** 32; + +end Interfaces; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/wordops.adb --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/wordops.adb Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,9 @@ +package body Wordops is + + function Rotate(I : Rotate_Amount; W : Word) return Word + is + begin + return Interfaces.Rotate_Left (W, I); + end Rotate; + +end Wordops; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/RIPEMD-160/wordops.ads --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/wordops.ads Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,19 @@ + +with Interfaces; +--# inherit Interfaces; + +package WordOps is + + subtype Word is Interfaces.Unsigned_32; + + subtype Rotate_Amount is Integer range 0..15; + + --# function rotate_left(I : Rotate_Amount; W : Word) return Word; + + function Rotate(I : Rotate_Amount; W : Word) return Word; + --# return rotate_left(I, W); + --# accept W, 3, "Expecting this warning"; + pragma Inline (Rotate); + --# end accept; + +end Wordops; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/ROOT.ML Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,16 @@ +use_thys + ["Gcd/Greatest_Common_Divisor", + + "Liseq/Longest_Increasing_Subsequence", + + "RIPEMD-160/F", + "RIPEMD-160/Hash", + "RIPEMD-160/K_L", + "RIPEMD-160/K_R", + "RIPEMD-160/R_L", + "RIPEMD-160/Round", + "RIPEMD-160/R_R", + "RIPEMD-160/S_L", + "RIPEMD-160/S_R", + + "Sqrt/Sqrt"]; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/Sqrt/Sqrt.adb --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/Sqrt/Sqrt.adb Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,16 @@ +package body Sqrt is + + function Isqrt(N: Natural) return Natural + is + R: Natural; + begin + R := 0; + loop + --# assert R * R <= N; + exit when N - R * R < 2 * R + 1; + R := R + 1; + end loop; + return R; + end Isqrt; + +end Sqrt; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/Sqrt/Sqrt.ads --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/Sqrt/Sqrt.ads Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,6 @@ +package Sqrt is + + function Isqrt(N: Natural) return Natural; + --# return R => R * R <= N and (R + 1) * (R + 1) > N; + +end Sqrt; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/Sqrt/Sqrt.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/Sqrt/Sqrt.thy Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,28 @@ +(* Title: HOL/SPARK/Examples/Sqrt/Sqrt.thy + Author: Stefan Berghofer + Copyright: secunet Security Networks AG +*) + +theory Sqrt +imports SPARK +begin + +spark_open "sqrt/isqrt.siv" + +spark_vc function_isqrt_4 +proof - + from `0 \ r` have "(r = 0 \ r = 1 \ r = 2) \ 2 < r" by auto + then show "2 * r \ 2147483646" + proof + assume "2 < r" + then have "0 < r" by simp + with `2 < r` have "2 * r < r * r" by (rule mult_strict_right_mono) + with `r * r \ n` and `n \ 2147483647` show ?thesis + by simp + qed auto + then show "2 * r \ 2147483647" by simp +qed + +spark_end + +end diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/Sqrt/sqrt/isqrt.fdl --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/Sqrt/sqrt/isqrt.fdl Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,29 @@ + {*******************************************************} + {FDL Declarations} + {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039} + {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.} + {*******************************************************} + + + {DATE : 29-NOV-2010 14:30:17.95} + + {function Sqrt.Isqrt} + + +title function isqrt; + + function round__(real) : integer; + const natural__base__first : integer = pending; + const natural__base__last : integer = pending; + const integer__base__first : integer = pending; + const integer__base__last : integer = pending; + const natural__first : integer = pending; + const natural__last : integer = pending; + const natural__size : integer = pending; + const integer__first : integer = pending; + const integer__last : integer = pending; + const integer__size : integer = pending; + var n : integer; + var r : integer; + +end; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/Sqrt/sqrt/isqrt.rls --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/Sqrt/sqrt/isqrt.rls Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,27 @@ + /*********************************************************/ + /*Proof Rule Declarations*/ + /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/ + /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/ + /*********************************************************/ + + + /*DATE : 29-NOV-2010 14:30:17.95*/ + + /*function Sqrt.Isqrt*/ + + +rule_family isqrt_rules: + X requires [X:any] & + X <= Y requires [X:ire, Y:ire] & + X >= Y requires [X:ire, Y:ire]. + +isqrt_rules(1): integer__size >= 0 may_be_deduced. +isqrt_rules(2): integer__first may_be_replaced_by -2147483648. +isqrt_rules(3): integer__last may_be_replaced_by 2147483647. +isqrt_rules(4): integer__base__first may_be_replaced_by -2147483648. +isqrt_rules(5): integer__base__last may_be_replaced_by 2147483647. +isqrt_rules(6): natural__size >= 0 may_be_deduced. +isqrt_rules(7): natural__first may_be_replaced_by 0. +isqrt_rules(8): natural__last may_be_replaced_by 2147483647. +isqrt_rules(9): natural__base__first may_be_replaced_by -2147483648. +isqrt_rules(10): natural__base__last may_be_replaced_by 2147483647. diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Examples/Sqrt/sqrt/isqrt.siv --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Examples/Sqrt/sqrt/isqrt.siv Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,64 @@ +***************************************************************************** + Semantic Analysis of SPARK Text + Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 + Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. +***************************************************************************** + + +CREATED 29-NOV-2010, 14:30:17 SIMPLIFIED 29-NOV-2010, 14:30:18 + +SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039 +Copyright (C) 2010 Altran Praxis Limited, Bath, U.K. + +function Sqrt.Isqrt + + + + +For path(s) from start to run-time check associated with statement of line 7: + +function_isqrt_1. +*** true . /* all conclusions proved */ + + +For path(s) from start to assertion of line 9: + +function_isqrt_2. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 9 to assertion of line 9: + +function_isqrt_3. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 9 to run-time check associated with + statement of line 10: + +function_isqrt_4. +H1: r * r <= n . +H2: n >= 0 . +H3: n <= 2147483647 . +H4: r >= 0 . +H5: r <= 2147483647 . +H6: integer__size >= 0 . +H7: natural__size >= 0 . + -> +C1: 2 * r <= 2147483646 . +C2: 2 * r <= 2147483647 . + + +For path(s) from assertion of line 9 to run-time check associated with + statement of line 11: + +function_isqrt_5. +*** true . /* all conclusions proved */ + + +For path(s) from assertion of line 9 to finish: + +function_isqrt_6. +*** true . /* all conclusions proved */ + + diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/ROOT.ML Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,1 @@ +use_thys ["SPARK"]; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/SPARK.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/SPARK.thy Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,278 @@ +(* Title: HOL/SPARK/SPARK.thy + Author: Stefan Berghofer + Copyright: secunet Security Networks AG + +Declaration of proof functions for SPARK/Ada verification environment. +*) + +theory SPARK +imports SPARK_Setup +begin + +text {* Bitwise logical operators *} + +spark_proof_functions + bit__and (integer, integer) : integer = "op AND" + bit__or (integer, integer) : integer = "op OR" + bit__xor (integer, integer) : integer = "op XOR" + +lemma AND_lower [simp]: + fixes x :: int and y :: int + assumes "0 \ x" + shows "0 \ x AND y" + using assms +proof (induct x arbitrary: y rule: bin_induct) + case (3 bin bit) + show ?case + proof (cases y rule: bin_exhaust) + case (1 bin' bit') + from 3 have "0 \ bin" + by (simp add: Bit_def bitval_def split add: bit.split_asm) + then have "0 \ bin AND bin'" by (rule 3) + with 1 show ?thesis + by simp (simp add: Bit_def bitval_def split add: bit.split) + qed +next + case 2 + then show ?case by (simp only: Min_def) +qed simp + +lemma OR_lower [simp]: + fixes x :: int and y :: int + assumes "0 \ x" "0 \ y" + shows "0 \ x OR y" + using assms +proof (induct x arbitrary: y rule: bin_induct) + case (3 bin bit) + show ?case + proof (cases y rule: bin_exhaust) + case (1 bin' bit') + from 3 have "0 \ bin" + by (simp add: Bit_def bitval_def split add: bit.split_asm) + moreover from 1 3 have "0 \ bin'" + by (simp add: Bit_def bitval_def split add: bit.split_asm) + ultimately have "0 \ bin OR bin'" by (rule 3) + with 1 show ?thesis + by simp (simp add: Bit_def bitval_def split add: bit.split) + qed +qed simp_all + +lemma XOR_lower [simp]: + fixes x :: int and y :: int + assumes "0 \ x" "0 \ y" + shows "0 \ x XOR y" + using assms +proof (induct x arbitrary: y rule: bin_induct) + case (3 bin bit) + show ?case + proof (cases y rule: bin_exhaust) + case (1 bin' bit') + from 3 have "0 \ bin" + by (simp add: Bit_def bitval_def split add: bit.split_asm) + moreover from 1 3 have "0 \ bin'" + by (simp add: Bit_def bitval_def split add: bit.split_asm) + ultimately have "0 \ bin XOR bin'" by (rule 3) + with 1 show ?thesis + by simp (simp add: Bit_def bitval_def split add: bit.split) + qed +next + case 2 + then show ?case by (simp only: Min_def) +qed simp + +lemma AND_upper1 [simp]: + fixes x :: int and y :: int + assumes "0 \ x" + shows "x AND y \ x" + using assms +proof (induct x arbitrary: y rule: bin_induct) + case (3 bin bit) + show ?case + proof (cases y rule: bin_exhaust) + case (1 bin' bit') + from 3 have "0 \ bin" + by (simp add: Bit_def bitval_def split add: bit.split_asm) + then have "bin AND bin' \ bin" by (rule 3) + with 1 show ?thesis + by simp (simp add: Bit_def bitval_def split add: bit.split) + qed +next + case 2 + then show ?case by (simp only: Min_def) +qed simp + +lemmas AND_upper1' [simp] = order_trans [OF AND_upper1] +lemmas AND_upper1'' [simp] = order_le_less_trans [OF AND_upper1] + +lemma AND_upper2 [simp]: + fixes x :: int and y :: int + assumes "0 \ y" + shows "x AND y \ y" + using assms +proof (induct y arbitrary: x rule: bin_induct) + case (3 bin bit) + show ?case + proof (cases x rule: bin_exhaust) + case (1 bin' bit') + from 3 have "0 \ bin" + by (simp add: Bit_def bitval_def split add: bit.split_asm) + then have "bin' AND bin \ bin" by (rule 3) + with 1 show ?thesis + by simp (simp add: Bit_def bitval_def split add: bit.split) + qed +next + case 2 + then show ?case by (simp only: Min_def) +qed simp + +lemmas AND_upper2' [simp] = order_trans [OF AND_upper2] +lemmas AND_upper2'' [simp] = order_le_less_trans [OF AND_upper2] + +lemma OR_upper: + fixes x :: int and y :: int + assumes "0 \ x" "x < 2 ^ n" "y < 2 ^ n" + shows "x OR y < 2 ^ n" + using assms +proof (induct x arbitrary: y n rule: bin_induct) + case (3 bin bit) + show ?case + proof (cases y rule: bin_exhaust) + case (1 bin' bit') + show ?thesis + proof (cases n) + case 0 + with 3 have "bin BIT bit = 0" by simp + then have "bin = 0" "bit = 0" + by (auto simp add: Bit_def bitval_def split add: bit.split_asm) arith + then show ?thesis using 0 1 `y < 2 ^ n` + by simp (simp add: Bit0_def int_or_Pls [unfolded Pls_def]) + next + case (Suc m) + from 3 have "0 \ bin" + by (simp add: Bit_def bitval_def split add: bit.split_asm) + moreover from 3 Suc have "bin < 2 ^ m" + by (simp add: Bit_def bitval_def split add: bit.split_asm) + moreover from 1 3 Suc have "bin' < 2 ^ m" + by (simp add: Bit_def bitval_def split add: bit.split_asm) + ultimately have "bin OR bin' < 2 ^ m" by (rule 3) + with 1 Suc show ?thesis + by simp (simp add: Bit_def bitval_def split add: bit.split) + qed + qed +qed simp_all + +lemmas [simp] = + OR_upper [of _ 8, simplified zle_diff1_eq [symmetric], simplified] + OR_upper [of _ 8, simplified] + OR_upper [of _ 16, simplified zle_diff1_eq [symmetric], simplified] + OR_upper [of _ 16, simplified] + OR_upper [of _ 32, simplified zle_diff1_eq [symmetric], simplified] + OR_upper [of _ 32, simplified] + OR_upper [of _ 64, simplified zle_diff1_eq [symmetric], simplified] + OR_upper [of _ 64, simplified] + +lemma XOR_upper: + fixes x :: int and y :: int + assumes "0 \ x" "x < 2 ^ n" "y < 2 ^ n" + shows "x XOR y < 2 ^ n" + using assms +proof (induct x arbitrary: y n rule: bin_induct) + case (3 bin bit) + show ?case + proof (cases y rule: bin_exhaust) + case (1 bin' bit') + show ?thesis + proof (cases n) + case 0 + with 3 have "bin BIT bit = 0" by simp + then have "bin = 0" "bit = 0" + by (auto simp add: Bit_def bitval_def split add: bit.split_asm) arith + then show ?thesis using 0 1 `y < 2 ^ n` + by simp (simp add: Bit0_def int_xor_Pls [unfolded Pls_def]) + next + case (Suc m) + from 3 have "0 \ bin" + by (simp add: Bit_def bitval_def split add: bit.split_asm) + moreover from 3 Suc have "bin < 2 ^ m" + by (simp add: Bit_def bitval_def split add: bit.split_asm) + moreover from 1 3 Suc have "bin' < 2 ^ m" + by (simp add: Bit_def bitval_def split add: bit.split_asm) + ultimately have "bin XOR bin' < 2 ^ m" by (rule 3) + with 1 Suc show ?thesis + by simp (simp add: Bit_def bitval_def split add: bit.split) + qed + qed +next + case 2 + then show ?case by (simp only: Min_def) +qed simp + +lemmas [simp] = + XOR_upper [of _ 8, simplified zle_diff1_eq [symmetric], simplified] + XOR_upper [of _ 8, simplified] + XOR_upper [of _ 16, simplified zle_diff1_eq [symmetric], simplified] + XOR_upper [of _ 16, simplified] + XOR_upper [of _ 32, simplified zle_diff1_eq [symmetric], simplified] + XOR_upper [of _ 32, simplified] + XOR_upper [of _ 64, simplified zle_diff1_eq [symmetric], simplified] + XOR_upper [of _ 64, simplified] + +lemma bit_not_spark_eq: + "NOT (word_of_int x :: ('a::len0) word) = + word_of_int (2 ^ len_of TYPE('a) - 1 - x)" +proof - + have "word_of_int x + NOT (word_of_int x) = + word_of_int x + (word_of_int (2 ^ len_of TYPE('a) - 1 - x)::'a word)" + by (simp only: bwsimps bin_add_not Min_def) + (simp add: word_of_int_hom_syms word_of_int_2p_len) + then show ?thesis by (rule add_left_imp_eq) +qed + +lemmas [simp] = + bit_not_spark_eq [where 'a=8, simplified] + bit_not_spark_eq [where 'a=16, simplified] + bit_not_spark_eq [where 'a=32, simplified] + bit_not_spark_eq [where 'a=64, simplified] + +lemma power_BIT: "2 ^ (Suc n) - 1 = (2 ^ n - 1) BIT 1" + unfolding Bit_B1 + by (induct n) simp_all + +lemma mod_BIT: + "bin BIT bit mod 2 ^ Suc n = (bin mod 2 ^ n) BIT bit" +proof - + have "bin mod 2 ^ n < 2 ^ n" by simp + then have "bin mod 2 ^ n \ 2 ^ n - 1" by simp + then have "2 * (bin mod 2 ^ n) \ 2 * (2 ^ n - 1)" + by (rule mult_left_mono) simp + then have "2 * (bin mod 2 ^ n) + 1 < 2 * 2 ^ n" by simp + then show ?thesis + by (auto simp add: Bit_def bitval_def mod_mult_mult1 mod_add_left_eq [of "2 * bin"] + mod_pos_pos_trivial split add: bit.split) +qed + +lemma AND_mod: + fixes x :: int + shows "x AND 2 ^ n - 1 = x mod 2 ^ n" +proof (induct x arbitrary: n rule: bin_induct) + case 1 + then show ?case + by simp (simp add: Pls_def) +next + case 2 + then show ?case + by (simp, simp only: Min_def, simp add: m1mod2k) +next + case (3 bin bit) + show ?case + proof (cases n) + case 0 + then show ?thesis by (simp add: int_and_extra_simps [unfolded Pls_def]) + next + case (Suc m) + with 3 show ?thesis + by (simp only: power_BIT mod_BIT int_and_Bits) simp + qed +qed + +end diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/SPARK_Setup.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/SPARK_Setup.thy Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,212 @@ +(* Title: HOL/SPARK/SPARK_Setup.thy + Author: Stefan Berghofer + Copyright: secunet Security Networks AG + +Setup for SPARK/Ada verification environment. +*) + +theory SPARK_Setup +imports Word +uses + "Tools/fdl_lexer.ML" + "Tools/fdl_parser.ML" + ("Tools/spark_vcs.ML") + ("Tools/spark_commands.ML") +begin + +text {* +SPARK versions of div and mod, see section 4.4.1.1 of SPARK Proof Manual +*} + +definition sdiv :: "int \ int \ int" (infixl "sdiv" 70) where + "a sdiv b = + (if 0 \ a then + if 0 \ b then a div b + else - (a div - b) + else + if 0 \ b then - (- a div b) + else - a div - b)" + +definition smod :: "int \ int \ int" (infixl "smod" 70) where + "a smod b = a - ((a sdiv b) * b)" + +lemma sdiv_minus_dividend: "- a sdiv b = - (a sdiv b)" + by (simp add: sdiv_def) + +lemma sdiv_minus_divisor: "a sdiv - b = - (a sdiv b)" + by (simp add: sdiv_def) + +lemma smod_minus_dividend: "- a smod b = - (a smod b)" + by (simp add: smod_def sdiv_minus_dividend) + +lemma smod_minus_divisor: "a smod - b = a smod b" + by (simp add: smod_def sdiv_minus_divisor) + +text {* +Correspondence between HOL's and SPARK's versions of div and mod +*} + +lemma sdiv_pos_pos: "0 \ a \ 0 \ b \ a sdiv b = a div b" + by (simp add: sdiv_def) + +lemma sdiv_pos_neg: "0 \ a \ b < 0 \ a sdiv b = - (a div - b)" + by (simp add: sdiv_def) + +lemma sdiv_neg_pos: "a < 0 \ 0 \ b \ a sdiv b = - (- a div b)" + by (simp add: sdiv_def) + +lemma sdiv_neg_neg: "a < 0 \ b < 0 \ a sdiv b = - a div - b" + by (simp add: sdiv_def) + +lemma smod_pos_pos: "0 \ a \ 0 \ b \ a smod b = a mod b" + by (simp add: smod_def sdiv_pos_pos zmod_zdiv_equality') + +lemma smod_pos_neg: "0 \ a \ b < 0 \ a smod b = a mod - b" + by (simp add: smod_def sdiv_pos_neg zmod_zdiv_equality') + +lemma smod_neg_pos: "a < 0 \ 0 \ b \ a smod b = - (- a mod b)" + by (simp add: smod_def sdiv_neg_pos zmod_zdiv_equality') + +lemma smod_neg_neg: "a < 0 \ b < 0 \ a smod b = - (- a mod - b)" + by (simp add: smod_def sdiv_neg_neg zmod_zdiv_equality') + + +text {* +Updating a function at a set of points. Useful for building arrays. +*} + +definition fun_upds :: "('a \ 'b) \ 'a set \ 'b \ 'a \ 'b" where + "fun_upds f xs y z = (if z \ xs then y else f z)" + +syntax + "_updsbind" :: "['a, 'a] => updbind" ("(2_ [:=]/ _)") + +translations + "f(xs[:=]y)" == "CONST fun_upds f xs y" + +lemma fun_upds_in [simp]: "z \ xs \ (f(xs [:=] y)) z = y" + by (simp add: fun_upds_def) + +lemma fun_upds_notin [simp]: "z \ xs \ (f(xs [:=] y)) z = f z" + by (simp add: fun_upds_def) + +lemma upds_singleton [simp]: "f({x} [:=] y) = f(x := y)" + by (simp add: fun_eq_iff) + + +text {* Enumeration types *} + +class enum = ord + finite + + fixes pos :: "'a \ int" + assumes range_pos: "range pos = {0.. y) = (pos x \ pos y)" +begin + +definition "val = inv pos" + +definition "succ x = val (pos x + 1)" + +definition "pred x = val (pos x - 1)" + +lemma inj_pos: "inj pos" + using finite_UNIV + by (rule eq_card_imp_inj_on) (simp add: range_pos) + +lemma val_pos: "val (pos x) = x" + unfolding val_def using inj_pos + by (rule inv_f_f) + +lemma pos_val: "z \ range pos \ pos (val z) = z" + unfolding val_def + by (rule f_inv_into_f) + +subclass linorder +proof + fix x::'a and y show "(x < y) = (x \ y \ \ y \ x)" + by (simp add: less_pos less_eq_pos less_le_not_le) +next + fix x::'a show "x \ x" by (simp add: less_eq_pos) +next + fix x::'a and y z assume "x \ y" and "y \ z" + then show "x \ z" by (simp add: less_eq_pos) +next + fix x::'a and y assume "x \ y" and "y \ x" + with inj_pos show "x = y" + by (auto dest: injD simp add: less_eq_pos) +next + fix x::'a and y show "x \ y \ y \ x" + by (simp add: less_eq_pos linear) +qed + +definition "first_el = val 0" + +definition "last_el = val (int (card (UNIV::'a set)) - 1)" + +lemma first_el_smallest: "first_el \ x" +proof - + have "pos x \ range pos" by (rule rangeI) + then have "pos (val 0) \ pos x" + by (simp add: range_pos pos_val) + then show ?thesis by (simp add: first_el_def less_eq_pos) +qed + +lemma last_el_greatest: "x \ last_el" +proof - + have "pos x \ range pos" by (rule rangeI) + then have "pos x \ pos (val (int (card (UNIV::'a set)) - 1))" + by (simp add: range_pos pos_val) + then show ?thesis by (simp add: last_el_def less_eq_pos) +qed + +lemma pos_succ: + assumes "x \ last_el" + shows "pos (succ x) = pos x + 1" +proof - + have "x \ last_el" by (rule last_el_greatest) + with assms have "x < last_el" by simp + then have "pos x < pos last_el" + by (simp add: less_pos) + with rangeI [of pos x] + have "pos x + 1 \ range pos" + by (simp add: range_pos last_el_def pos_val) + then show ?thesis + by (simp add: succ_def pos_val) +qed + +lemma pos_pred: + assumes "x \ first_el" + shows "pos (pred x) = pos x - 1" +proof - + have "first_el \ x" by (rule first_el_smallest) + with assms have "first_el < x" by simp + then have "pos first_el < pos x" + by (simp add: less_pos) + with rangeI [of pos x] + have "pos x - 1 \ range pos" + by (simp add: range_pos first_el_def pos_val) + then show ?thesis + by (simp add: pred_def pos_val) +qed + +lemma succ_val: "x \ range pos \ succ (val x) = val (x + 1)" + by (simp add: succ_def pos_val) + +lemma pred_val: "x \ range pos \ pred (val x) = val (x - 1)" + by (simp add: pred_def pos_val) + +end + +lemma interval_expand: + "x < y \ (z::int) \ {x.. z \ {x+1.. 'a * chars) -> (chars -> T * chars) -> + Position.T -> string -> 'a * T list + val position_of: T -> Position.T + val pos_of: T -> string + val is_eof: T -> bool + val stopper: T Scan.stopper + val kind_of: T -> kind + val content_of: T -> string + val unparse: T -> string + val is_proper: T -> bool + val is_digit: string -> bool + val c_comment: chars -> T * chars + val curly_comment: chars -> T * chars + val percent_comment: chars -> T * chars + val vcg_header: chars -> (banner * (date * time) option) * chars + val siv_header: chars -> + (banner * (date * time) * (date * time) * (string * string)) * chars +end; + +structure Fdl_Lexer: FDL_LEXER = +struct + +(** tokens **) + +datatype kind = Keyword | Ident | Long_Ident | Traceability | Number | Comment | EOF; + +datatype T = Token of kind * string * Position.T; + +fun make_token k xs = Token (k, implode (map fst xs), + case xs of [] => Position.none | (_, p) :: _ => p); + +fun kind_of (Token (k, _, _)) = k; + +fun is_proper (Token (Comment, _, _)) = false + | is_proper _ = true; + +fun content_of (Token (_, s, _)) = s; + +fun unparse (Token (Traceability, s, _)) = "For " ^ s ^ ":" + | unparse (Token (_, s, _)) = s; + +fun position_of (Token (_, _, pos)) = pos; + +val pos_of = Position.str_of o position_of; + +fun is_eof (Token (EOF, _, _)) = true + | is_eof _ = false; + +fun mk_eof pos = Token (EOF, "", pos); +val eof = mk_eof Position.none; + +val stopper = + Scan.stopper (fn [] => eof | toks => mk_eof (position_of (List.last toks))) is_eof; + +fun leq_token (Token (_, s, _), Token (_, s', _)) = size s <= size s'; + + +(** split up a string into a list of characters (with positions) **) + +type chars = (string * Position.T) list; + +fun is_char_eof ("", _) = true + | is_char_eof _ = false; + +val char_stopper = Scan.stopper (K ("", Position.none)) is_char_eof; + +fun symbol (x : string, _ : Position.T) = x; + +fun explode_pos s pos = fst (fold_map (fn x => fn pos => + ((x, pos), Position.advance x pos)) (raw_explode s) pos); + + +(** scanners **) + +val any = Scan.one (not o Scan.is_stopper char_stopper); + +fun prfx [] = Scan.succeed [] + | prfx (x :: xs) = Scan.one (equal x o symbol) ::: prfx xs; + +val $$$ = prfx o raw_explode; + +val lexicon = Scan.make_lexicon (map raw_explode + ["rule_family", + "title", + "For", + ":", + "[", + "]", + "(", + ")", + ",", + "&", + ";", + "=", + ".", + "..", + "requires", + "may_be_replaced_by", + "may_be_deduced", + "may_be_deduced_from", + "are_interchangeable", + "if", + "end", + "function", + "procedure", + "type", + "var", + "const", + "array", + "record", + ":=", + "of", + "**", + "*", + "/", + "div", + "mod", + "+", + "-", + "<>", + "<", + ">", + "<=", + ">=", + "<->", + "->", + "not", + "and", + "or", + "for_some", + "for_all", + "***", + "!!!", + "element", + "update", + "pending"]); + +fun keyword s = Scan.literal lexicon :|-- + (fn xs => if map symbol xs = raw_explode s then Scan.succeed xs else Scan.fail); + +fun is_digit x = "0" <= x andalso x <= "9"; +fun is_alpha x = "a" <= x andalso x <= "z" orelse "A" <= x andalso x <= "Z"; +val is_underscore = equal "_"; +val is_tilde = equal "~"; +val is_newline = equal "\n"; +val is_tab = equal "\t"; +val is_space = equal " "; +val is_whitespace = is_space orf is_tab orf is_newline; +val is_whitespace' = is_space orf is_tab; + +val number = Scan.many1 (is_digit o symbol); + +val identifier = + Scan.one (is_alpha o symbol) ::: + Scan.many + ((is_alpha orf is_digit orf is_underscore) o symbol) @@@ + Scan.optional (Scan.one (is_tilde o symbol) >> single) []; + +val long_identifier = + identifier @@@ (Scan.repeat1 ($$$ "." @@@ identifier) >> flat); + +val whitespace = Scan.many (is_whitespace o symbol); +val whitespace' = Scan.many (is_whitespace' o symbol); +val newline = Scan.one (is_newline o symbol); + +fun beginning n cs = + let + val drop_blanks = #1 o take_suffix is_whitespace; + val all_cs = drop_blanks cs; + val dots = if length all_cs > n then " ..." else ""; + in + (drop_blanks (take n all_cs) + |> map (fn c => if is_whitespace c then " " else c) + |> implode) ^ dots + end; + +fun !!! text scan = + let + fun get_pos [] = " (past end-of-text!)" + | get_pos ((_, pos) :: _) = Position.str_of pos; + + fun err (syms, msg) = + text ^ get_pos syms ^ " at " ^ beginning 10 (map symbol syms) ^ + (case msg of NONE => "" | SOME s => "\n" ^ s); + in Scan.!! err scan end; + +val any_line' = + Scan.many (not o (Scan.is_stopper char_stopper orf (is_newline o symbol))); + +val any_line = whitespace' |-- any_line' --| + newline >> (implode o map symbol); + +fun gen_comment a b = $$$ a |-- !!! "missing end of comment" + (Scan.repeat (Scan.unless ($$$ b) any) --| $$$ b) >> make_token Comment; + +val c_comment = gen_comment "/*" "*/"; +val curly_comment = gen_comment "{" "}"; + +val percent_comment = $$$ "%" |-- any_line' >> make_token Comment; + +fun repeatn 0 _ = Scan.succeed [] + | repeatn n p = Scan.one p ::: repeatn (n-1) p; + + +(** header of *.vcg and *.siv files (see simplifier/load_provenance.pro) **) + +type banner = string * string * string; +type date = string * string * string; +type time = string * string * string * string option; + +val asterisks = Scan.repeat1 (Scan.one (equal "*" o symbol)); + +fun alphan n = repeatn n (is_alpha o symbol) >> (implode o map symbol); +fun digitn n = repeatn n (is_digit o symbol) >> (implode o map symbol); + +val time = + digitn 2 --| $$$ ":" -- digitn 2 --| $$$ ":" -- digitn 2 -- + Scan.option ($$$ "." |-- digitn 2) >> + (fn (((hr, mi), s), ms) => (hr, mi, s, ms)); + +val date = + digitn 2 --| $$$ "-" -- alphan 3 --| $$$ "-" -- digitn 4 >> + (fn ((d, m), y) => (d, m, y)); + +val banner = + whitespace' |-- asterisks --| whitespace' --| newline :|-- (fn xs => + (any_line -- any_line -- any_line >> + (fn ((l1, l2), l3) => (l1, l2, l3))) --| + whitespace' --| prfx (map symbol xs) --| whitespace' --| newline); + +val vcg_header = banner -- Scan.option (whitespace |-- + $$$ "DATE :" |-- whitespace |-- date --| whitespace --| + Scan.option ($$$ "TIME :" -- whitespace) -- time); + +val siv_header = banner --| whitespace -- + ($$$ "CREATED" |-- whitespace |-- (date --| $$$ "," --| whitespace -- time)) --| + whitespace -- + ($$$ "SIMPLIFIED" |-- whitespace |-- (date --| $$$ "," --| whitespace -- time)) --| + newline --| newline -- (any_line -- any_line) >> + (fn (((b, c), s), ls) => (b, c, s, ls)); + + +(** the main tokenizer **) + +fun scan header comment = + !!! "bad header" header --| whitespace -- + Scan.repeat (Scan.unless (Scan.one is_char_eof) + (!!! "bad input" + ( comment + || (keyword "For" -- whitespace) |-- + Scan.repeat1 (Scan.unless (keyword ":") any) --| + keyword ":" >> make_token Traceability + || Scan.max leq_token + (Scan.literal lexicon >> make_token Keyword) + ( long_identifier >> make_token Long_Ident + || identifier >> make_token Ident) + || number >> make_token Number) --| + whitespace)); + +fun tokenize header comment pos s = + fst (Scan.finite char_stopper + (Scan.error (scan header comment)) (explode_pos s pos)); + +end; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Tools/fdl_parser.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Tools/fdl_parser.ML Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,384 @@ +(* Title: HOL/SPARK/Tools/fdl_parser.ML + Author: Stefan Berghofer + Copyright: secunet Security Networks AG + +Parser for fdl files. +*) + +signature FDL_PARSER = +sig + datatype expr = + Ident of string + | Number of int + | Quantifier of string * string list * string * expr + | Funct of string * expr list + | Element of expr * expr list + | Update of expr * expr list * expr + | Record of string * (string * expr) list + | Array of string * expr option * + ((expr * expr option) list list * expr) list; + + datatype fdl_type = + Basic_Type of string + | Enum_Type of string list + | Array_Type of string list * string + | Record_Type of (string list * string) list + | Pending_Type; + + datatype fdl_rule = + Inference_Rule of expr list * expr + | Substitution_Rule of expr list * expr * expr; + + type rules = + ((string * int) * fdl_rule) list * + (string * (expr * (string * string) list) list) list; + + type vcs = (string * (string * + ((string * expr) list * (string * expr) list)) list) list; + + type 'a tab = 'a Symtab.table * (string * 'a) list; + + val lookup: 'a tab -> string -> 'a option; + val update: string * 'a -> 'a tab -> 'a tab; + val items: 'a tab -> (string * 'a) list; + + type decls = + {types: fdl_type tab, + vars: string tab, + consts: string tab, + funs: (string list * string) tab}; + + val parse_vcs: (Fdl_Lexer.chars -> 'a * Fdl_Lexer.chars) -> Position.T -> + string -> 'a * ((string * string) * vcs); + + val parse_declarations: Position.T -> string -> (string * string) * decls; + + val parse_rules: Position.T -> string -> rules; +end; + +structure Fdl_Parser: FDL_PARSER = +struct + +(** error handling **) + +fun beginning n cs = + let val dots = if length cs > n then " ..." else ""; + in + space_implode " " (take n cs) ^ dots + end; + +fun !!! scan = + let + fun get_pos [] = " (past end-of-file!)" + | get_pos (tok :: _) = Fdl_Lexer.pos_of tok; + + fun err (syms, msg) = + "Syntax error" ^ get_pos syms ^ " at " ^ + beginning 10 (map Fdl_Lexer.unparse syms) ^ + (case msg of NONE => "" | SOME s => "\n" ^ s ^ " expected"); + in Scan.!! err scan end; + +fun parse_all p = + Scan.repeat (Scan.unless (Scan.one Fdl_Lexer.is_eof) (!!! p)); + + +(** parsers **) + +fun group s p = p || Scan.fail_with (K s); + +fun $$$ s = group s + (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Keyword andalso + Fdl_Lexer.content_of t = s) >> K s); + +val identifier = group "identifier" + (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Ident) >> + Fdl_Lexer.content_of); + +val long_identifier = group "long identifier" + (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Long_Ident) >> + Fdl_Lexer.content_of); + +fun the_identifier s = group s + (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Ident andalso + Fdl_Lexer.content_of t = s) >> K s); + +fun prfx_identifier g s = group g + (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Ident andalso + can (unprefix s) (Fdl_Lexer.content_of t)) >> + (unprefix s o Fdl_Lexer.content_of)); + +val mk_identifier = prfx_identifier "identifier \"mk__...\"" "mk__"; +val hyp_identifier = prfx_identifier "hypothesis identifer" "H"; +val concl_identifier = prfx_identifier "conclusion identifier" "C"; + +val number = group "number" + (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Number) >> + (the o Int.fromString o Fdl_Lexer.content_of)); + +val traceability = group "traceability information" + (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Traceability) >> + Fdl_Lexer.content_of); + +fun enum1 sep scan = scan ::: Scan.repeat ($$$ sep |-- !!! scan); +fun enum sep scan = enum1 sep scan || Scan.succeed []; + +fun list1 scan = enum1 "," scan; +fun list scan = enum "," scan; + + +(* expressions, see section 4.4 of SPARK Proof Manual *) + +datatype expr = + Ident of string + | Number of int + | Quantifier of string * string list * string * expr + | Funct of string * expr list + | Element of expr * expr list + | Update of expr * expr list * expr + | Record of string * (string * expr) list + | Array of string * expr option * + ((expr * expr option) list list * expr) list; + +fun unop (f, x) = Funct (f, [x]); + +fun binop p q = p :|-- (fn x => Scan.optional + (q -- !!! p >> (fn (f, y) => Funct (f, [x, y]))) x); + +(* left-associative *) +fun binops opp argp = + argp -- Scan.repeat (opp -- !!! argp) >> (fn (x, fys) => + fold (fn (f, y) => fn x => Funct (f, [x, y])) fys x); + +(* right-associative *) +fun binops' f p = enum1 f p >> foldr1 (fn (x, y) => Funct (f, [x, y])); + +val multiplying_operator = $$$ "*" || $$$ "/" || $$$ "div" || $$$ "mod"; + +val adding_operator = $$$ "+" || $$$ "-"; + +val relational_operator = + $$$ "=" || $$$ "<>" + || $$$ "<" || $$$ ">" + || $$$ "<="|| $$$ ">="; + +val quantification_kind = $$$ "for_all" || $$$ "for_some"; + +val quantification_generator = + list1 identifier --| $$$ ":" -- identifier; + +fun expression xs = group "expression" + (binop disjunction ($$$ "->" || $$$ "<->")) xs + +and disjunction xs = binops' "or" conjunction xs + +and conjunction xs = binops' "and" negation xs + +and negation xs = + ( $$$ "not" -- !!! relation >> unop + || relation) xs + +and relation xs = binop sum relational_operator xs + +and sum xs = binops adding_operator term xs + +and term xs = binops multiplying_operator factor xs + +and factor xs = + ( $$$ "+" |-- !!! primary + || $$$ "-" -- !!! primary >> unop + || binop primary ($$$ "**")) xs + +and primary xs = group "primary" + ( number >> Number + || $$$ "(" |-- !!! (expression --| $$$ ")") + || quantified_expression + || function_designator + || identifier >> Ident) xs + +and quantified_expression xs = (quantification_kind -- + !!! ($$$ "(" |-- quantification_generator --| $$$ "," -- + expression --| $$$ ")") >> (fn (q, ((xs, T), e)) => + Quantifier (q, xs, T, e))) xs + +and function_designator xs = + ( mk_identifier --| $$$ "(" :|-- + (fn s => record_args s || array_args s) --| $$$ ")" + || $$$ "element" |-- !!! ($$$ "(" |-- expression --| $$$ "," --| $$$ "[" -- + list1 expression --| $$$ "]" --| $$$ ")") >> Element + || $$$ "update" |-- !!! ($$$ "(" |-- expression --| $$$ "," --| $$$ "[" -- + list1 expression --| $$$ "]" --| $$$ "," -- expression --| $$$ ")") >> + (fn ((A, xs), x) => Update (A, xs, x)) + || identifier --| $$$ "(" -- !!! (list1 expression --| $$$ ")") >> Funct) xs + +and record_args s xs = + (list1 (identifier --| $$$ ":=" -- !!! expression) >> (pair s #> Record)) xs + +and array_args s xs = + ( expression -- Scan.optional ($$$ "," |-- !!! array_associations) [] >> + (fn (default, assocs) => Array (s, SOME default, assocs)) + || array_associations >> (fn assocs => Array (s, NONE, assocs))) xs + +and array_associations xs = + (list1 (enum1 "&" ($$$ "[" |-- + !!! (list1 (expression -- Scan.option ($$$ ".." |-- !!! expression)) --| + $$$ "]")) --| $$$ ":=" -- expression)) xs; + + +(* verification conditions *) + +type vcs = (string * (string * + ((string * expr) list * (string * expr) list)) list) list; + +val vc = + identifier --| $$$ "." -- !!! + ( $$$ "***" |-- !!! (the_identifier "true" --| $$$ ".") >> + (Ident #> pair "1" #> single #> pair []) + || $$$ "!!!" |-- !!! (the_identifier "false" --| $$$ ".") >> + (Ident #> pair "1" #> single #> pair []) + || Scan.repeat1 (hyp_identifier -- + !!! ($$$ ":" |-- expression --| $$$ ".")) --| $$$ "->" -- + Scan.repeat1 (concl_identifier -- + !!! ($$$ ":" |-- expression --| $$$ "."))); + +val subprogram_kind = $$$ "function" || $$$ "procedure"; + +val vcs = + subprogram_kind -- (long_identifier || identifier) -- + parse_all (traceability -- !!! (Scan.repeat1 vc)); + +fun parse_vcs header pos s = + s |> + Fdl_Lexer.tokenize header Fdl_Lexer.c_comment pos ||> + filter Fdl_Lexer.is_proper ||> + Scan.finite Fdl_Lexer.stopper (Scan.error (!!! vcs)) ||> + fst; + + +(* fdl declarations, see section 4.3 of SPARK Proof Manual *) + +datatype fdl_type = + Basic_Type of string + | Enum_Type of string list + | Array_Type of string list * string + | Record_Type of (string list * string) list + | Pending_Type; + +(* also store items in a list to preserve order *) +type 'a tab = 'a Symtab.table * (string * 'a) list; + +fun lookup ((tab, _) : 'a tab) = Symtab.lookup tab; +fun update decl (tab, items) = (Symtab.update_new decl tab, decl :: items); +fun items ((_, items) : 'a tab) = rev items; + +type decls = + {types: fdl_type tab, + vars: string tab, + consts: string tab, + funs: (string list * string) tab}; + +val empty_decls : decls = + {types = (Symtab.empty, []), vars = (Symtab.empty, []), + consts = (Symtab.empty, []), funs = (Symtab.empty, [])}; + +fun add_type_decl decl {types, vars, consts, funs} = + {types = update decl types, + vars = vars, consts = consts, funs = funs} + handle Symtab.DUP s => error ("Duplicate type " ^ s); + +fun add_var_decl (vs, ty) {types, vars, consts, funs} = + {types = types, + vars = fold (update o rpair ty) vs vars, + consts = consts, funs = funs} + handle Symtab.DUP s => error ("Duplicate variable " ^ s); + +fun add_const_decl decl {types, vars, consts, funs} = + {types = types, vars = vars, + consts = update decl consts, + funs = funs} + handle Symtab.DUP s => error ("Duplicate constant " ^ s); + +fun add_fun_decl decl {types, vars, consts, funs} = + {types = types, vars = vars, consts = consts, + funs = update decl funs} + handle Symtab.DUP s => error ("Duplicate function " ^ s); + +val type_decl = $$$ "type" |-- !!! (identifier --| $$$ "=" -- + ( identifier >> Basic_Type + || $$$ "(" |-- !!! (list1 identifier --| $$$ ")") >> Enum_Type + || $$$ "array" |-- !!! ($$$ "[" |-- list1 identifier --| $$$ "]" --| + $$$ "of" -- identifier) >> Array_Type + || $$$ "record" |-- !!! (enum1 ";" + (list1 identifier -- !!! ($$$ ":" |-- identifier)) --| + $$$ "end") >> Record_Type + || $$$ "pending" >> K Pending_Type)) >> add_type_decl; + +val const_decl = $$$ "const" |-- !!! (identifier --| $$$ ":" -- identifier --| + $$$ "=" --| $$$ "pending") >> add_const_decl; + +val var_decl = $$$ "var" |-- !!! (list1 identifier --| $$$ ":" -- identifier) >> + add_var_decl; + +val fun_decl = $$$ "function" |-- !!! (identifier -- + (Scan.optional ($$$ "(" |-- !!! (list1 identifier --| $$$ ")")) [] --| + $$$ ":" -- identifier)) >> add_fun_decl; + +val declarations = + $$$ "title" |-- subprogram_kind -- identifier --| $$$ ";" -- + (Scan.repeat ((type_decl || const_decl || var_decl || fun_decl) --| + !!! ($$$ ";")) >> (fn ds => apply ds empty_decls)) --| + $$$ "end" --| $$$ ";" + +fun parse_declarations pos s = + s |> + Fdl_Lexer.tokenize (Scan.succeed ()) Fdl_Lexer.curly_comment pos |> + snd |> filter Fdl_Lexer.is_proper |> + Scan.finite Fdl_Lexer.stopper (Scan.error (!!! declarations)) |> + fst; + + +(* rules, see section 5 of SPADE Proof Checker Rules Manual *) + +datatype fdl_rule = + Inference_Rule of expr list * expr + | Substitution_Rule of expr list * expr * expr; + +type rules = + ((string * int) * fdl_rule) list * + (string * (expr * (string * string) list) list) list; + +val condition_list = $$$ "[" |-- list expression --| $$$ "]"; +val if_condition_list = $$$ "if" |-- !!! condition_list; + +val rule = + identifier -- !!! ($$$ "(" |-- number --| $$$ ")" --| $$$ ":" -- + (expression :|-- (fn e => + $$$ "may_be_deduced" >> K (Inference_Rule ([], e)) + || $$$ "may_be_deduced_from" |-- + !!! condition_list >> (Inference_Rule o rpair e) + || $$$ "may_be_replaced_by" |-- !!! (expression -- + Scan.optional if_condition_list []) >> (fn (e', cs) => + Substitution_Rule (cs, e, e')) + || $$$ "&" |-- !!! (expression --| $$$ "are_interchangeable" -- + Scan.optional if_condition_list []) >> (fn (e', cs) => + Substitution_Rule (cs, e, e')))) --| $$$ ".") >> + (fn (id, (n, rl)) => ((id, n), rl)); + +val rule_family = + $$$ "rule_family" |-- identifier --| $$$ ":" -- + enum1 "&" (expression -- !!! ($$$ "requires" |-- $$$ "[" |-- + list (identifier -- !!! ($$$ ":" |-- identifier)) --| $$$ "]")) --| + $$$ "."; + +val rules = + parse_all (rule >> (apfst o cons) || rule_family >> (apsnd o cons)) >> + (fn rls => apply (rev rls) ([], [])); + +fun parse_rules pos s = + s |> + Fdl_Lexer.tokenize (Scan.succeed ()) + (Fdl_Lexer.c_comment || Fdl_Lexer.percent_comment) pos |> + snd |> filter Fdl_Lexer.is_proper |> + Scan.finite Fdl_Lexer.stopper (Scan.error (!!! rules)) |> + fst; + +end; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Tools/spark_commands.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Tools/spark_commands.ML Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,150 @@ +(* Title: HOL/SPARK/Tools/spark_commands.ML + Author: Stefan Berghofer + Copyright: secunet Security Networks AG + +Isar commands for handling SPARK/Ada verification conditions. +*) + +signature SPARK_COMMANDS = +sig + val setup: theory -> theory +end + +structure SPARK_Commands: SPARK_COMMANDS = +struct + +fun read f path = f (Position.file (Path.implode path)) (File.read path); + +fun spark_open vc_name thy = + let + val (vc_path, _) = Thy_Load.check_file + [Thy_Load.master_directory thy] (Path.explode vc_name); + val (base, header) = (case Path.split_ext vc_path of + (base, "vcg") => (base, Fdl_Lexer.vcg_header >> K ()) + | (base, "siv") => (base, Fdl_Lexer.siv_header >> K ()) + | _ => error "File name must end with .vcg or .siv"); + val fdl_path = Path.ext "fdl" base; + val rls_path = Path.ext "rls" base; + in + SPARK_VCs.set_vcs + (snd (read Fdl_Parser.parse_declarations fdl_path)) + (read Fdl_Parser.parse_rules rls_path) + (snd (snd (read (Fdl_Parser.parse_vcs header) vc_path))) + base thy + end; + +fun add_proof_fun_cmd pf thy = + let val ctxt = ProofContext.init_global thy + in SPARK_VCs.add_proof_fun + (fn optT => Syntax.parse_term ctxt #> + the_default I (Option.map Type.constraint optT) #> + Syntax.check_term ctxt) pf thy + end; + +fun get_vc thy vc_name = + (case SPARK_VCs.lookup_vc thy vc_name of + SOME (ctxt, (_, proved, ctxt', stmt)) => + if proved then + error ("The verification condition " ^ + quote vc_name ^ " has already been proved.") + else (ctxt @ [ctxt'], stmt) + | NONE => error ("There is no verification condition " ^ + quote vc_name ^ ".")); + +fun prove_vc vc_name lthy = + let + val thy = ProofContext.theory_of lthy; + val (ctxt, stmt) = get_vc thy vc_name + in + Specification.theorem Thm.theoremK NONE + (K (Local_Theory.background_theory (SPARK_VCs.mark_proved vc_name))) + (Binding.name vc_name, []) ctxt stmt true lthy + end; + +fun string_of_status false = "(unproved)" + | string_of_status true = "(proved)"; + +fun show_status (p, f) = Toplevel.no_timing o Toplevel.keep (fn state => + let + val thy = Toplevel.theory_of state; + + val (context, defs, vcs) = SPARK_VCs.get_vcs thy; + + val vcs' = AList.coalesce (op =) (map_filter + (fn (name, (trace, status, ctxt, stmt)) => + if p status then + SOME (trace, (name, status, ctxt, stmt)) + else NONE) vcs); + + val ctxt = state |> + Toplevel.theory_of |> + ProofContext.init_global |> + Context.proof_map (fold Element.init context) + in + [Pretty.str "Context:", + Pretty.chunks (maps (Element.pretty_ctxt ctxt) context), + + Pretty.str "Definitions:", + Pretty.chunks (map (fn (bdg, th) => Pretty.block + [Pretty.str (Binding.str_of bdg ^ ":"), + Pretty.brk 1, + Display.pretty_thm ctxt th]) + defs), + + Pretty.str "Verification conditions:", + Pretty.chunks2 (maps (fn (trace, vcs'') => + Pretty.str trace :: + map (fn (name, status, context', stmt) => + Pretty.big_list (name ^ " " ^ f status) + (Element.pretty_ctxt ctxt context' @ + Element.pretty_stmt ctxt stmt)) vcs'') vcs')] |> + Pretty.chunks2 |> Pretty.writeln + end); + +val _ = + Outer_Syntax.command "spark_open" + "open a new SPARK environment and load a SPARK-generated .vcg or .siv file" + Keyword.thy_decl + (Parse.name >> (Toplevel.theory o spark_open)); + +val pfun_type = Scan.option + (Args.parens (Parse.list1 Parse.name) --| Args.colon -- Parse.name); + +val _ = + Outer_Syntax.command "spark_proof_functions" + "associate SPARK proof functions with terms" + Keyword.thy_decl + (Scan.repeat1 (Parse.name -- (pfun_type --| Args.$$$ "=" -- Parse.term)) >> + (Toplevel.theory o fold add_proof_fun_cmd)); + +val _ = + Outer_Syntax.command "spark_vc" + "enter into proof mode for a specific verification condition" + Keyword.thy_goal + (Parse.name >> (fn name => + (Toplevel.print o Toplevel.local_theory_to_proof NONE (prove_vc name)))); + +val _ = + Outer_Syntax.improper_command "spark_status" + "show the name and state of all loaded verification conditions" + Keyword.diag + (Scan.optional + (Args.parens + ( Args.$$$ "proved" >> K (I, K "") + || Args.$$$ "unproved" >> K (not, K ""))) + (K true, string_of_status) >> show_status); + +val _ = + Outer_Syntax.command "spark_end" + "close the current SPARK environment" + Keyword.thy_decl + (Scan.succeed (Toplevel.theory SPARK_VCs.close)); + +val setup = Theory.at_end (fn thy => + let + val _ = SPARK_VCs.is_closed thy + orelse error ("Found the end of the theory, " ^ + "but the last SPARK environment is still open.") + in NONE end); + +end; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/SPARK/Tools/spark_vcs.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SPARK/Tools/spark_vcs.ML Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,870 @@ +(* Title: HOL/SPARK/Tools/spark_vcs.ML + Author: Stefan Berghofer + Copyright: secunet Security Networks AG + +Store for verification conditions generated by SPARK/Ada. +*) + +signature SPARK_VCS = +sig + val set_vcs: Fdl_Parser.decls -> Fdl_Parser.rules -> Fdl_Parser.vcs -> + Path.T -> theory -> theory + val add_proof_fun: (typ option -> 'a -> term) -> + string * ((string list * string) option * 'a) -> + theory -> theory + val lookup_vc: theory -> string -> (Element.context_i list * + (string * bool * Element.context_i * Element.statement_i)) option + val get_vcs: theory -> + Element.context_i list * (binding * thm) list * + (string * (string * bool * Element.context_i * Element.statement_i)) list + val mark_proved: string -> theory -> theory + val close: theory -> theory + val is_closed: theory -> bool +end; + +structure SPARK_VCs: SPARK_VCS = +struct + +open Fdl_Parser; + + +(** utilities **) + +fun mk_unop s t = + let val T = fastype_of t + in Const (s, T --> T) $ t end; + +fun mk_times (t, u) = + let + val setT = fastype_of t; + val T = HOLogic.dest_setT setT; + val U = HOLogic.dest_setT (fastype_of u) + in + Const (@{const_name Sigma}, setT --> (T --> HOLogic.mk_setT U) --> + HOLogic.mk_setT (HOLogic.mk_prodT (T, U))) $ t $ Abs ("", T, u) + end; + +fun mk_type _ "integer" = HOLogic.intT + | mk_type _ "boolean" = HOLogic.boolT + | mk_type thy ty = Syntax.check_typ (ProofContext.init_global thy) + (Type (Sign.full_name thy (Binding.name ty), [])); + +val booleanN = "boolean"; +val integerN = "integer"; + +fun mk_qual_name thy s s' = + Sign.full_name thy (Binding.qualify true s (Binding.name s')); + +fun define_overloaded (def_name, eq) lthy = + let + val ((c, _), rhs) = eq |> Syntax.check_term lthy |> + Logic.dest_equals |>> dest_Free; + val ((_, (_, thm)), lthy') = Local_Theory.define + ((Binding.name c, NoSyn), ((Binding.name def_name, []), rhs)) lthy + val ctxt_thy = ProofContext.init_global (ProofContext.theory_of lthy'); + val thm' = singleton (ProofContext.export lthy' ctxt_thy) thm + in (thm', lthy') end; + +fun strip_underscores s = + strip_underscores (unsuffix "_" s) handle Fail _ => s; + +fun strip_tilde s = + unsuffix "~" s ^ "_init" handle Fail _ => s; + +val mangle_name = strip_underscores #> strip_tilde; + +fun mk_variables thy xs ty (tab, ctxt) = + let + val T = mk_type thy ty; + val (ys, ctxt') = Name.variants (map mangle_name xs) ctxt; + val zs = map (Free o rpair T) ys; + in (zs, (fold (Symtab.update o apsnd (rpair ty)) (xs ~~ zs) tab, ctxt')) end; + + +(** generate properties of enumeration types **) + +fun add_enum_type tyname els (tab, ctxt) thy = + let + val tyb = Binding.name tyname; + val tyname' = Sign.full_name thy tyb; + val T = Type (tyname', []); + val case_name = mk_qual_name thy tyname (tyname ^ "_case"); + val cs = map (fn s => Const (mk_qual_name thy tyname s, T)) els; + val k = length els; + val p = Const (@{const_name pos}, T --> HOLogic.intT); + val v = Const (@{const_name val}, HOLogic.intT --> T); + val card = Const (@{const_name card}, + HOLogic.mk_setT T --> HOLogic.natT) $ HOLogic.mk_UNIV T; + + fun mk_binrel_def s f = Logic.mk_equals + (Const (s, T --> T --> HOLogic.boolT), + Abs ("x", T, Abs ("y", T, + Const (s, HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $ + (f $ Bound 1) $ (f $ Bound 0)))); + + val (((def1, def2), def3), lthy) = thy |> + Datatype.add_datatype {strict = true, quiet = true} [tyname] + [([], tyb, NoSyn, + map (fn s => (Binding.name s, [], NoSyn)) els)] |> snd |> + + Class.instantiation ([tyname'], [], @{sort enum}) |> + + define_overloaded ("pos_" ^ tyname ^ "_def", Logic.mk_equals + (p, + list_comb (Const (case_name, replicate k HOLogic.intT @ + [T] ---> HOLogic.intT), + map (HOLogic.mk_number HOLogic.intT) (0 upto k - 1)))) ||>> + + define_overloaded ("less_eq_" ^ tyname ^ "_def", + mk_binrel_def @{const_name less_eq} p) ||>> + define_overloaded ("less_" ^ tyname ^ "_def", + mk_binrel_def @{const_name less} p); + + val UNIV_eq = Goal.prove lthy [] [] + (HOLogic.mk_Trueprop (HOLogic.mk_eq + (HOLogic.mk_UNIV T, HOLogic.mk_set T cs))) + (fn _ => + rtac @{thm subset_antisym} 1 THEN + rtac @{thm subsetI} 1 THEN + Datatype_Aux.exh_tac (K (#exhaust (Datatype_Data.the_info + (ProofContext.theory_of lthy) tyname'))) 1 THEN + ALLGOALS (asm_full_simp_tac (simpset_of lthy))); + + val finite_UNIV = Goal.prove lthy [] [] + (HOLogic.mk_Trueprop (Const (@{const_name finite}, + HOLogic.mk_setT T --> HOLogic.boolT) $ HOLogic.mk_UNIV T)) + (fn _ => simp_tac (simpset_of lthy addsimps [UNIV_eq]) 1); + + val card_UNIV = Goal.prove lthy [] [] + (HOLogic.mk_Trueprop (HOLogic.mk_eq + (card, HOLogic.mk_number HOLogic.natT k))) + (fn _ => simp_tac (simpset_of lthy addsimps [UNIV_eq]) 1); + + val range_pos = Goal.prove lthy [] [] + (HOLogic.mk_Trueprop (HOLogic.mk_eq + (Const (@{const_name image}, (T --> HOLogic.intT) --> + HOLogic.mk_setT T --> HOLogic.mk_setT HOLogic.intT) $ + p $ HOLogic.mk_UNIV T, + Const (@{const_name atLeastLessThan}, HOLogic.intT --> + HOLogic.intT --> HOLogic.mk_setT HOLogic.intT) $ + HOLogic.mk_number HOLogic.intT 0 $ + (@{term int} $ card)))) + (fn _ => + simp_tac (simpset_of lthy addsimps [card_UNIV]) 1 THEN + simp_tac (simpset_of lthy addsimps [UNIV_eq, def1]) 1 THEN + rtac @{thm subset_antisym} 1 THEN + simp_tac (simpset_of lthy) 1 THEN + rtac @{thm subsetI} 1 THEN + asm_full_simp_tac (simpset_of lthy addsimps @{thms interval_expand} + delsimps @{thms atLeastLessThan_iff}) 1); + + val lthy' = + Class.prove_instantiation_instance (fn _ => + Class.intro_classes_tac [] THEN + rtac finite_UNIV 1 THEN + rtac range_pos 1 THEN + simp_tac (HOL_basic_ss addsimps [def3]) 1 THEN + simp_tac (HOL_basic_ss addsimps [def2]) 1) lthy; + + val (pos_eqs, val_eqs) = split_list (map_index (fn (i, c) => + let + val n = HOLogic.mk_number HOLogic.intT i; + val th = Goal.prove lthy' [] [] + (HOLogic.mk_Trueprop (HOLogic.mk_eq (p $ c, n))) + (fn _ => simp_tac (simpset_of lthy' addsimps [def1]) 1); + val th' = Goal.prove lthy' [] [] + (HOLogic.mk_Trueprop (HOLogic.mk_eq (v $ n, c))) + (fn _ => + rtac (@{thm inj_pos} RS @{thm injD}) 1 THEN + simp_tac (simpset_of lthy' addsimps + [@{thm pos_val}, range_pos, card_UNIV, th]) 1) + in (th, th') end) cs); + + val first_el = Goal.prove lthy' [] [] + (HOLogic.mk_Trueprop (HOLogic.mk_eq + (Const (@{const_name first_el}, T), hd cs))) + (fn _ => simp_tac (simpset_of lthy' addsimps + [@{thm first_el_def}, hd val_eqs]) 1); + + val last_el = Goal.prove lthy' [] [] + (HOLogic.mk_Trueprop (HOLogic.mk_eq + (Const (@{const_name last_el}, T), List.last cs))) + (fn _ => simp_tac (simpset_of lthy' addsimps + [@{thm last_el_def}, List.last val_eqs, card_UNIV]) 1); + + val simp_att = [Attrib.internal (K Simplifier.simp_add)] + + in + ((fold (Symtab.update_new o apsnd (rpair tyname)) (els ~~ cs) tab, + fold Name.declare els ctxt), + lthy' |> + Local_Theory.note + ((Binding.name (tyname ^ "_card_UNIV"), simp_att), [card_UNIV]) ||>> + Local_Theory.note + ((Binding.name (tyname ^ "_pos"), simp_att), pos_eqs) ||>> + Local_Theory.note + ((Binding.name (tyname ^ "_val"), simp_att), val_eqs) ||>> + Local_Theory.note + ((Binding.name (tyname ^ "_first_el"), simp_att), [first_el]) ||>> + Local_Theory.note + ((Binding.name (tyname ^ "_last_el"), simp_att), [last_el]) |> snd |> + Local_Theory.exit_global) + end; + + +fun add_type_def (s, Basic_Type ty) (ids, thy) = + (ids, + Typedecl.abbrev_global (Binding.name s, [], NoSyn) + (mk_type thy ty) thy |> snd) + + | add_type_def (s, Enum_Type els) (ids, thy) = add_enum_type s els ids thy + + | add_type_def (s, Array_Type (argtys, resty)) (ids, thy) = + (ids, + Typedecl.abbrev_global (Binding.name s, [], NoSyn) + (foldr1 HOLogic.mk_prodT (map (mk_type thy) argtys) --> + mk_type thy resty) thy |> snd) + + | add_type_def (s, Record_Type fldtys) (ids, thy) = + (ids, + Record.add_record true ([], Binding.name s) NONE + (maps (fn (flds, ty) => + let val T = mk_type thy ty + in map (fn fld => (Binding.name fld, T, NoSyn)) flds + end) fldtys) thy) + + | add_type_def (s, Pending_Type) (ids, thy) = + (ids, Typedecl.typedecl_global (Binding.name s, [], NoSyn) thy |> snd); + + +fun term_of_expr thy types funs pfuns = + let + fun tm_of vs (Funct ("->", [e, e'])) = + (HOLogic.mk_imp (fst (tm_of vs e), fst (tm_of vs e')), booleanN) + + | tm_of vs (Funct ("<->", [e, e'])) = + (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e')), booleanN) + + | tm_of vs (Funct ("or", [e, e'])) = + (HOLogic.mk_disj (fst (tm_of vs e), fst (tm_of vs e')), booleanN) + + | tm_of vs (Funct ("and", [e, e'])) = + (HOLogic.mk_conj (fst (tm_of vs e), fst (tm_of vs e')), booleanN) + + | tm_of vs (Funct ("not", [e])) = + (HOLogic.mk_not (fst (tm_of vs e)), booleanN) + + | tm_of vs (Funct ("=", [e, e'])) = + (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e')), booleanN) + + | tm_of vs (Funct ("<>", [e, e'])) = (HOLogic.mk_not + (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e'))), booleanN) + + | tm_of vs (Funct ("<", [e, e'])) = (HOLogic.mk_binrel @{const_name less} + (fst (tm_of vs e), fst (tm_of vs e')), booleanN) + + | tm_of vs (Funct (">", [e, e'])) = (HOLogic.mk_binrel @{const_name less} + (fst (tm_of vs e'), fst (tm_of vs e)), booleanN) + + | tm_of vs (Funct ("<=", [e, e'])) = (HOLogic.mk_binrel @{const_name less_eq} + (fst (tm_of vs e), fst (tm_of vs e')), booleanN) + + | tm_of vs (Funct (">=", [e, e'])) = (HOLogic.mk_binrel @{const_name less_eq} + (fst (tm_of vs e'), fst (tm_of vs e)), booleanN) + + | tm_of vs (Funct ("+", [e, e'])) = (HOLogic.mk_binop @{const_name plus} + (fst (tm_of vs e), fst (tm_of vs e')), integerN) + + | tm_of vs (Funct ("-", [e, e'])) = (HOLogic.mk_binop @{const_name minus} + (fst (tm_of vs e), fst (tm_of vs e')), integerN) + + | tm_of vs (Funct ("*", [e, e'])) = (HOLogic.mk_binop @{const_name times} + (fst (tm_of vs e), fst (tm_of vs e')), integerN) + + | tm_of vs (Funct ("/", [e, e'])) = (HOLogic.mk_binop @{const_name divide} + (fst (tm_of vs e), fst (tm_of vs e')), integerN) + + | tm_of vs (Funct ("div", [e, e'])) = (HOLogic.mk_binop @{const_name sdiv} + (fst (tm_of vs e), fst (tm_of vs e')), integerN) + + | tm_of vs (Funct ("mod", [e, e'])) = (HOLogic.mk_binop @{const_name smod} + (fst (tm_of vs e), fst (tm_of vs e')), integerN) + + | tm_of vs (Funct ("-", [e])) = + (mk_unop @{const_name uminus} (fst (tm_of vs e)), integerN) + + | tm_of vs (Funct ("**", [e, e'])) = + (Const (@{const_name power}, HOLogic.intT --> HOLogic.natT --> + HOLogic.intT) $ fst (tm_of vs e) $ + (@{const nat} $ fst (tm_of vs e')), integerN) + + | tm_of (tab, _) (Ident s) = + (case Symtab.lookup tab s of + SOME t_ty => t_ty + | NONE => error ("Undeclared identifier " ^ s)) + + | tm_of _ (Number i) = (HOLogic.mk_number HOLogic.intT i, integerN) + + | tm_of vs (Quantifier (s, xs, ty, e)) = + let + val (ys, vs') = mk_variables thy xs ty vs; + val q = (case s of + "for_all" => HOLogic.mk_all + | "for_some" => HOLogic.mk_exists) + in + (fold_rev (fn Free (x, T) => fn t => q (x, T, t)) + ys (fst (tm_of vs' e)), + booleanN) + end + + | tm_of vs (Funct (s, es)) = + + (* record field selection *) + (case try (unprefix "fld_") s of + SOME fname => (case es of + [e] => + let val (t, rcdty) = tm_of vs e + in case lookup types rcdty of + SOME (Record_Type fldtys) => + (case get_first (fn (flds, fldty) => + if member (op =) flds fname then SOME fldty + else NONE) fldtys of + SOME fldty => + (Const (mk_qual_name thy rcdty fname, + mk_type thy rcdty --> mk_type thy fldty) $ t, + fldty) + | NONE => error ("Record " ^ rcdty ^ + " has no field named " ^ fname)) + | _ => error (rcdty ^ " is not a record type") + end + | _ => error ("Function " ^ s ^ " expects one argument")) + | NONE => + + (* record field update *) + (case try (unprefix "upf_") s of + SOME fname => (case es of + [e, e'] => + let + val (t, rcdty) = tm_of vs e; + val rT = mk_type thy rcdty; + val (u, fldty) = tm_of vs e'; + val fT = mk_type thy fldty + in case lookup types rcdty of + SOME (Record_Type fldtys) => + (case get_first (fn (flds, fldty) => + if member (op =) flds fname then SOME fldty + else NONE) fldtys of + SOME fldty' => + if fldty = fldty' then + (Const (mk_qual_name thy rcdty (fname ^ "_update"), + (fT --> fT) --> rT --> rT) $ + Abs ("x", fT, u) $ t, + rcdty) + else error ("Type " ^ fldty ^ + " does not match type " ^ fldty' ^ " of field " ^ + fname) + | NONE => error ("Record " ^ rcdty ^ + " has no field named " ^ fname)) + | _ => error (rcdty ^ " is not a record type") + end + | _ => error ("Function " ^ s ^ " expects two arguments")) + | NONE => + + (* enumeration type to integer *) + (case try (unsuffix "__pos") s of + SOME tyname => (case es of + [e] => (Const (@{const_name pos}, + mk_type thy tyname --> HOLogic.intT) $ fst (tm_of vs e), integerN) + | _ => error ("Function " ^ s ^ " expects one argument")) + | NONE => + + (* integer to enumeration type *) + (case try (unsuffix "__val") s of + SOME tyname => (case es of + [e] => (Const (@{const_name val}, + HOLogic.intT --> mk_type thy tyname) $ fst (tm_of vs e), tyname) + | _ => error ("Function " ^ s ^ " expects one argument")) + | NONE => + + (* successor / predecessor of enumeration type element *) + if s = "succ" orelse s = "pred" then (case es of + [e] => + let + val (t, tyname) = tm_of vs e; + val T = mk_type thy tyname + in (Const + (if s = "succ" then @{const_name succ} + else @{const_name pred}, T --> T) $ t, tyname) + end + | _ => error ("Function " ^ s ^ " expects one argument")) + + (* user-defined proof function *) + else + (case Symtab.lookup pfuns s of + SOME (SOME (_, resty), t) => + (list_comb (t, map (fst o tm_of vs) es), resty) + | _ => error ("Undeclared proof function " ^ s)))))) + + | tm_of vs (Element (e, es)) = + let val (t, ty) = tm_of vs e + in case lookup types ty of + SOME (Array_Type (_, elty)) => + (t $ foldr1 HOLogic.mk_prod (map (fst o tm_of vs) es), elty) + | _ => error (ty ^ " is not an array type") + end + + | tm_of vs (Update (e, es, e')) = + let val (t, ty) = tm_of vs e + in case lookup types ty of + SOME (Array_Type (idxtys, elty)) => + let + val T = foldr1 HOLogic.mk_prodT (map (mk_type thy) idxtys); + val U = mk_type thy elty; + val fT = T --> U + in + (Const (@{const_name fun_upd}, fT --> T --> U --> fT) $ + t $ foldr1 HOLogic.mk_prod (map (fst o tm_of vs) es) $ + fst (tm_of vs e'), + ty) + end + | _ => error (ty ^ " is not an array type") + end + + | tm_of vs (Record (s, flds)) = + (case lookup types s of + SOME (Record_Type fldtys) => + let + val flds' = map (apsnd (tm_of vs)) flds; + val fnames = maps fst fldtys; + val fnames' = map fst flds; + val (fvals, ftys) = split_list (map (fn s' => + case AList.lookup (op =) flds' s' of + SOME fval_ty => fval_ty + | NONE => error ("Field " ^ s' ^ " missing in record " ^ s)) + fnames); + val _ = (case subtract (op =) fnames fnames' of + [] => () + | xs => error ("Extra field(s) " ^ commas xs ^ + " in record " ^ s)); + val _ = (case duplicates (op =) fnames' of + [] => () + | xs => error ("Duplicate field(s) " ^ commas xs ^ + " in record " ^ s)) + in + (list_comb + (Const (mk_qual_name thy s (s ^ "_ext"), + map (mk_type thy) ftys @ [HOLogic.unitT] ---> + mk_type thy s), + fvals @ [HOLogic.unit]), + s) + end + | _ => error (s ^ " is not a record type")) + + | tm_of vs (Array (s, default, assocs)) = + (case lookup types s of + SOME (Array_Type (idxtys, elty)) => + let + val Ts = map (mk_type thy) idxtys; + val T = foldr1 HOLogic.mk_prodT Ts; + val U = mk_type thy elty; + fun mk_idx' T (e, NONE) = HOLogic.mk_set T [fst (tm_of vs e)] + | mk_idx' T (e, SOME e') = Const (@{const_name atLeastAtMost}, + T --> T --> HOLogic.mk_setT T) $ + fst (tm_of vs e) $ fst (tm_of vs e'); + fun mk_idx idx = + if length Ts <> length idx then + error ("Arity mismatch in construction of array " ^ s) + else foldr1 mk_times (map2 mk_idx' Ts idx); + fun mk_upd (idxs, e) t = + if length idxs = 1 andalso forall (is_none o snd) (hd idxs) + then + Const (@{const_name fun_upd}, (T --> U) --> + T --> U --> T --> U) $ t $ + foldl1 HOLogic.mk_prod + (map (fst o tm_of vs o fst) (hd idxs)) $ + fst (tm_of vs e) + else + Const (@{const_name fun_upds}, (T --> U) --> + HOLogic.mk_setT T --> U --> T --> U) $ t $ + foldl1 (HOLogic.mk_binop @{const_name sup}) + (map mk_idx idxs) $ + fst (tm_of vs e) + in + (fold mk_upd assocs + (case default of + SOME e => Abs ("x", T, fst (tm_of vs e)) + | NONE => Const (@{const_name undefined}, T --> U)), + s) + end + | _ => error (s ^ " is not an array type")) + + in tm_of end; + + +fun term_of_rule thy types funs pfuns ids rule = + let val tm_of = fst o term_of_expr thy types funs pfuns ids + in case rule of + Inference_Rule (es, e) => Logic.list_implies + (map (HOLogic.mk_Trueprop o tm_of) es, HOLogic.mk_Trueprop (tm_of e)) + | Substitution_Rule (es, e, e') => Logic.list_implies + (map (HOLogic.mk_Trueprop o tm_of) es, + HOLogic.mk_Trueprop (HOLogic.mk_eq (tm_of e, tm_of e'))) + end; + + +val builtin = Symtab.make (map (rpair ()) + ["->", "<->", "or", "and", "not", "=", "<>", "<", ">", "<=", ">=", + "+", "-", "*", "/", "div", "mod", "**"]); + +fun complex_expr (Number _) = false + | complex_expr (Ident _) = false + | complex_expr (Funct (s, es)) = + not (Symtab.defined builtin s) orelse exists complex_expr es + | complex_expr (Quantifier (_, _, _, e)) = complex_expr e + | complex_expr _ = true; + +fun complex_rule (Inference_Rule (es, e)) = + complex_expr e orelse exists complex_expr es + | complex_rule (Substitution_Rule (es, e, e')) = + complex_expr e orelse complex_expr e' orelse + exists complex_expr es; + +val is_pfun = + Symtab.defined builtin orf + can (unprefix "fld_") orf can (unprefix "upf_") orf + can (unsuffix "__pos") orf can (unsuffix "__val") orf + equal "succ" orf equal "pred"; + +fun fold_opt f = the_default I o Option.map f; +fun fold_pair f g (x, y) = f x #> g y; + +fun fold_expr f g (Funct (s, es)) = f s #> fold (fold_expr f g) es + | fold_expr f g (Ident s) = g s + | fold_expr f g (Number _) = I + | fold_expr f g (Quantifier (_, _, _, e)) = fold_expr f g e + | fold_expr f g (Element (e, es)) = + fold_expr f g e #> fold (fold_expr f g) es + | fold_expr f g (Update (e, es, e')) = + fold_expr f g e #> fold (fold_expr f g) es #> fold_expr f g e' + | fold_expr f g (Record (_, flds)) = fold (fold_expr f g o snd) flds + | fold_expr f g (Array (_, default, assocs)) = + fold_opt (fold_expr f g) default #> + fold (fold_pair + (fold (fold (fold_pair + (fold_expr f g) (fold_opt (fold_expr f g))))) + (fold_expr f g)) assocs; + +val add_expr_pfuns = fold_expr + (fn s => if is_pfun s then I else insert (op =) s) (K I); + +val add_expr_idents = fold_expr (K I) (insert (op =)); + +fun pfun_type thy (argtys, resty) = + map (mk_type thy) argtys ---> mk_type thy resty; + +fun check_pfun_type thy s t optty1 optty2 = + let + val T = fastype_of t; + fun check ty = + let val U = pfun_type thy ty + in + T = U orelse + error ("Type\n" ^ + Syntax.string_of_typ_global thy T ^ + "\nof function " ^ + Syntax.string_of_term_global thy t ^ + " associated with proof function " ^ s ^ + "\ndoes not match declared type\n" ^ + Syntax.string_of_typ_global thy U) + end + in (Option.map check optty1; Option.map check optty2; ()) end; + +fun upd_option x y = if is_some x then x else y; + +fun check_pfuns_types thy funs = + Symtab.map (fn s => fn (optty, t) => + let val optty' = lookup funs s + in + (check_pfun_type thy s t optty optty'; + (NONE |> upd_option optty |> upd_option optty', t)) + end); + + +(** the VC store **) + +fun err_unfinished () = error "An unfinished SPARK environment is still open." + +fun err_vcs names = error (Pretty.string_of + (Pretty.big_list "The following verification conditions have not been proved:" + (map Pretty.str names))) + +val strip_number = pairself implode o take_suffix Fdl_Lexer.is_digit o raw_explode; + +val name_ord = prod_ord string_ord (option_ord int_ord) o + pairself (strip_number ##> Int.fromString); + +structure VCtab = Table(type key = string val ord = name_ord); + +structure VCs = Theory_Data +( + type T = + {pfuns: ((string list * string) option * term) Symtab.table, + env: + {ctxt: Element.context_i list, + defs: (binding * thm) list, + types: fdl_type tab, + funs: (string list * string) tab, + ids: (term * string) Symtab.table * Name.context, + proving: bool, + vcs: (string * bool * + (string * expr) list * (string * expr) list) VCtab.table, + path: Path.T} option} + val empty : T = {pfuns = Symtab.empty, env = NONE} + val extend = I + fun merge ({pfuns = pfuns1, env = NONE}, {pfuns = pfuns2, env = NONE}) = + {pfuns = Symtab.merge (eq_pair (op =) (op aconv)) (pfuns1, pfuns2), + env = NONE} + | merge _ = err_unfinished () +) + +fun set_env (env as {funs, ...}) thy = VCs.map (fn + {pfuns, env = NONE} => + {pfuns = check_pfuns_types thy funs pfuns, env = SOME env} + | _ => err_unfinished ()) thy; + +fun mk_pat s = (case Int.fromString s of + SOME i => [HOLogic.mk_Trueprop (Var (("C", i), HOLogic.boolT))] + | NONE => error ("Bad conclusion identifier: C" ^ s)); + +fun mk_vc thy types funs pfuns ids (tr, proved, ps, cs) = + let val prop_of = + HOLogic.mk_Trueprop o fst o term_of_expr thy types funs pfuns ids + in + (tr, proved, + Element.Assumes (map (fn (s', e) => + ((Binding.name ("H" ^ s'), []), [(prop_of e, [])])) ps), + Element.Shows (map (fn (s', e) => + (Attrib.empty_binding, [(prop_of e, mk_pat s')])) cs)) + end; + +fun fold_vcs f vcs = + VCtab.fold (fn (_, (_, _, ps, cs)) => fold f ps #> fold f cs) vcs; + +fun pfuns_of_vcs pfuns vcs = + fold_vcs (add_expr_pfuns o snd) vcs [] |> + filter_out (Symtab.defined pfuns); + +fun declare_missing_pfuns thy funs pfuns vcs (tab, ctxt) = + let + val (fs, (tys, Ts)) = + pfuns_of_vcs pfuns vcs |> + map_filter (fn s => lookup funs s |> + Option.map (fn ty => (s, (SOME ty, pfun_type thy ty)))) |> + split_list ||> split_list; + val (fs', ctxt') = Name.variants fs ctxt + in + (fold Symtab.update_new (fs ~~ (tys ~~ map Free (fs' ~~ Ts))) pfuns, + Element.Fixes (map2 (fn s => fn T => + (Binding.name s, SOME T, NoSyn)) fs' Ts), + (tab, ctxt')) + end; + +fun add_proof_fun prep (s, (optty, raw_t)) thy = + VCs.map (fn + {env = SOME {proving = true, ...}, ...} => err_unfinished () + | {pfuns, env} => + let + val optty' = (case env of + SOME {funs, ...} => lookup funs s + | NONE => NONE); + val optty'' = NONE |> upd_option optty |> upd_option optty'; + val t = prep (Option.map (pfun_type thy) optty'') raw_t + in + (check_pfun_type thy s t optty optty'; + if is_some optty'' orelse is_none env then + {pfuns = Symtab.update_new (s, (optty'', t)) pfuns, + env = env} + handle Symtab.DUP _ => error ("Proof function " ^ s ^ + " already associated with function") + else error ("Undeclared proof function " ^ s)) + end) thy; + +val is_closed = is_none o #env o VCs.get; + +fun lookup_vc thy name = + (case VCs.get thy of + {env = SOME {vcs, types, funs, ids, ctxt, ...}, pfuns} => + (case VCtab.lookup vcs name of + SOME vc => + let val (pfuns', ctxt', ids') = + declare_missing_pfuns thy funs pfuns vcs ids + in SOME (ctxt @ [ctxt'], mk_vc thy types funs pfuns' ids' vc) end + | NONE => NONE) + | _ => NONE); + +fun get_vcs thy = (case VCs.get thy of + {env = SOME {vcs, types, funs, ids, ctxt, defs, ...}, pfuns} => + let val (pfuns', ctxt', ids') = + declare_missing_pfuns thy funs pfuns vcs ids + in + (ctxt @ [ctxt'], defs, + VCtab.dest vcs |> + map (apsnd (mk_vc thy types funs pfuns' ids'))) + end + | _ => ([], [], [])); + +fun mark_proved name = VCs.map (fn + {pfuns, env = SOME {ctxt, defs, types, funs, ids, vcs, path, ...}} => + {pfuns = pfuns, + env = SOME {ctxt = ctxt, defs = defs, + types = types, funs = funs, ids = ids, + proving = true, + vcs = VCtab.map_entry name (fn (trace, _, ps, cs) => + (trace, true, ps, cs)) vcs, + path = path}} + | x => x); + +fun close thy = VCs.map (fn + {pfuns, env = SOME {vcs, path, ...}} => + (case VCtab.fold_rev (fn (s, (_, p, _, _)) => + (if p then apfst else apsnd) (cons s)) vcs ([], []) of + (proved, []) => + (File.write (Path.ext "prv" path) + (concat (map (fn s => snd (strip_number s) ^ + " -- proved by " ^ Distribution.version ^ "\n") proved)); + {pfuns = pfuns, env = NONE}) + | (_, unproved) => err_vcs unproved) + | x => x) thy; + + +(** set up verification conditions **) + +fun partition_opt f = + let + fun part ys zs [] = (rev ys, rev zs) + | part ys zs (x :: xs) = (case f x of + SOME y => part (y :: ys) zs xs + | NONE => part ys (x :: zs) xs) + in part [] [] end; + +fun dest_def (id, (Substitution_Rule ([], Ident s, rhs))) = SOME (id, (s, rhs)) + | dest_def _ = NONE; + +fun mk_rulename (s, i) = Binding.name (s ^ string_of_int i); + +fun add_const (s, ty) ((tab, ctxt), thy) = + let + val T = mk_type thy ty; + val b = Binding.name s; + val c = Const (Sign.full_name thy b, T) + in + (c, + ((Symtab.update (s, (c, ty)) tab, Name.declare s ctxt), + Sign.add_consts_i [(b, T, NoSyn)] thy)) + end; + +fun add_def types funs pfuns consts (id, (s, e)) (ids as (tab, ctxt), thy) = + (case lookup consts s of + SOME ty => + let + val (t, ty') = term_of_expr thy types funs pfuns ids e; + val _ = ty = ty' orelse + error ("Declared type " ^ ty ^ " of " ^ s ^ + "\ndoes not match type " ^ ty' ^ " in definition"); + val id' = mk_rulename id; + val lthy = Named_Target.theory_init thy; + val ((t', (_, th)), lthy') = Specification.definition + (NONE, ((id', []), HOLogic.mk_Trueprop (HOLogic.mk_eq + (Free (s, mk_type thy ty), t)))) lthy; + val phi = ProofContext.export_morphism lthy' lthy + in + ((id', Morphism.thm phi th), + ((Symtab.update (s, (Morphism.term phi t', ty)) tab, + Name.declare s ctxt), + Local_Theory.exit_global lthy')) + end + | NONE => error ("Undeclared constant " ^ s)); + +fun add_var (s, ty) (ids, thy) = + let val ([Free p], ids') = mk_variables thy [s] ty ids + in (p, (ids', thy)) end; + +fun add_init_vars vcs (ids_thy as ((tab, _), _)) = + fold_map add_var + (map_filter + (fn s => case try (unsuffix "~") s of + SOME s' => (case Symtab.lookup tab s' of + SOME (_, ty) => SOME (s, ty) + | NONE => error ("Undeclared identifier " ^ s')) + | NONE => NONE) + (fold_vcs (add_expr_idents o snd) vcs [])) + ids_thy; + +fun is_trivial_vc ([], [(_, Ident "true")]) = true + | is_trivial_vc _ = false; + +fun rulenames rules = commas + (map (fn ((s, i), _) => s ^ "(" ^ string_of_int i ^ ")") rules); + +(* sort definitions according to their dependency *) +fun sort_defs _ _ [] sdefs = rev sdefs + | sort_defs pfuns consts defs sdefs = + (case find_first (fn (_, (_, e)) => + forall (Symtab.defined pfuns) (add_expr_pfuns e []) andalso + forall (fn id => + member (fn (s, (_, (s', _))) => s = s') sdefs id orelse + member (fn (s, (s', _)) => s = s') consts id) + (add_expr_idents e [])) defs of + SOME d => sort_defs pfuns consts + (remove (op =) d defs) (d :: sdefs) + | NONE => error ("Bad definitions: " ^ rulenames defs)); + +fun set_vcs ({types, vars, consts, funs} : decls) (rules, _) vcs path thy = + let + val {pfuns, ...} = VCs.get thy; + val (defs', rules') = partition_opt dest_def rules; + val consts' = + subtract (fn ((_, (s, _)), (s', _)) => s = s') defs' (items consts); + val defs = sort_defs pfuns consts' defs' []; + (* ignore all complex rules in rls files *) + val (rules'', other_rules) = + List.partition (complex_rule o snd) rules'; + val _ = if null rules'' then () + else warning ("Ignoring rules: " ^ rulenames rules''); + + val vcs' = VCtab.make (maps (fn (tr, vcs) => + map (fn (s, (ps, cs)) => (s, (tr, false, ps, cs))) + (filter_out (is_trivial_vc o snd) vcs)) vcs); + + val _ = (case filter_out (is_some o lookup funs) + (pfuns_of_vcs pfuns vcs') of + [] => () + | fs => error ("Undeclared proof function(s) " ^ commas fs)); + + val (((defs', vars''), ivars), (ids, thy')) = + ((Symtab.empty |> + Symtab.update ("false", (HOLogic.false_const, booleanN)) |> + Symtab.update ("true", (HOLogic.true_const, booleanN)), + Name.context), thy) |> + fold add_type_def (items types) |> + fold (snd oo add_const) consts' |> + fold_map (add_def types funs pfuns consts) defs ||>> + fold_map add_var (items vars) ||>> + add_init_vars vcs'; + + val ctxt = + [Element.Fixes (map (fn (s, T) => + (Binding.name s, SOME T, NoSyn)) (vars'' @ ivars)), + Element.Assumes (map (fn (id, rl) => + ((mk_rulename id, []), + [(term_of_rule thy' types funs pfuns ids rl, [])])) + other_rules), + Element.Notes (Thm.definitionK, + [((Binding.name "defns", []), map (rpair [] o single o snd) defs')])] + + in + set_env {ctxt = ctxt, defs = defs', types = types, funs = funs, + ids = ids, proving = false, vcs = vcs', path = path} thy' + end; + +end; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Statespace/ROOT.ML --- a/src/HOL/Statespace/ROOT.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Statespace/ROOT.ML Tue Jan 25 09:45:45 2011 +0100 @@ -1,1 +1,1 @@ -use_thys ["StateSpaceEx"]; \ No newline at end of file +use_thys ["StateSpaceEx"]; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Statespace/state_space.ML --- a/src/HOL/Statespace/state_space.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Statespace/state_space.ML Tue Jan 25 09:45:45 2011 +0100 @@ -145,20 +145,20 @@ fun prove_interpretation_in ctxt_tac (name, expr) thy = thy - |> Expression.sublocale_cmd name expr [] + |> Expression.sublocale_cmd I name expr [] |> Proof.global_terminal_proof (Method.Basic (fn ctxt => SIMPLE_METHOD (ctxt_tac ctxt)), NONE) |> ProofContext.theory_of fun add_locale name expr elems thy = thy - |> Expression.add_locale (Binding.name name) (Binding.name name) expr elems + |> Expression.add_locale I (Binding.name name) (Binding.name name) expr elems |> snd |> Local_Theory.exit; fun add_locale_cmd name expr elems thy = thy - |> Expression.add_locale_cmd (Binding.name name) Binding.empty expr elems + |> Expression.add_locale_cmd I (Binding.name name) Binding.empty expr elems |> snd |> Local_Theory.exit; @@ -349,7 +349,7 @@ fun add_declaration name decl thy = thy - |> Named_Target.init name + |> Named_Target.init I name |> (fn lthy => Local_Theory.declaration false (decl lthy) lthy) |> Local_Theory.exit_global; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/TLA/Buffer/Buffer.thy --- a/src/HOL/TLA/Buffer/Buffer.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/TLA/Buffer/Buffer.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,8 +1,5 @@ -(* - File: Buffer.thy - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich +(* Title: HOL/TLA/Buffer/Buffer.thy + Author: Stephan Merz, University of Munich *) header {* A simple FIFO buffer (synchronous communication, interleaving) *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/TLA/Buffer/DBuffer.thy --- a/src/HOL/TLA/Buffer/DBuffer.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/TLA/Buffer/DBuffer.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,8 +1,5 @@ -(* - File: DBuffer.thy - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich +(* Title: HOL/TLA/Buffer/DBuffer.thy + Author: Stephan Merz, University of Munich *) header {* Two FIFO buffers in a row, with interleaving assumption *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/TLA/Inc/Inc.thy --- a/src/HOL/TLA/Inc/Inc.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/TLA/Inc/Inc.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,8 +1,5 @@ -(* - File: TLA/Inc/Inc.thy - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich +(* Title: HOL/TLA/Inc/Inc.thy + Author: Stephan Merz, University of Munich *) header {* Lamport's "increment" example *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/TLA/Memory/MemClerk.thy --- a/src/HOL/TLA/Memory/MemClerk.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/TLA/Memory/MemClerk.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,8 +1,5 @@ -(* - File: MemClerk.thy - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich +(* Title: HOL/TLA/Memory/MemClerk.thy + Author: Stephan Merz, University of Munich *) header {* RPC-Memory example: specification of the memory clerk *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/TLA/Memory/MemClerkParameters.thy --- a/src/HOL/TLA/Memory/MemClerkParameters.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/TLA/Memory/MemClerkParameters.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,8 +1,5 @@ -(* - File: MemClerkParameters.thy - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich +(* Title: HOL/TLA/Memory/MemClerkParameters.thy + Author: Stephan Merz, University of Munich *) header {* RPC-Memory example: Parameters of the memory clerk *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/TLA/Memory/Memory.thy --- a/src/HOL/TLA/Memory/Memory.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/TLA/Memory/Memory.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,8 +1,5 @@ -(* - File: Memory.thy - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich +(* Title: HOL/TLA/Memory/Memory.thy + Author: Stephan Merz, University of Munich *) header {* RPC-Memory example: Memory specification *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/TLA/Memory/MemoryImplementation.thy --- a/src/HOL/TLA/Memory/MemoryImplementation.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/TLA/Memory/MemoryImplementation.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,8 +1,5 @@ -(* - File: MemoryImplementation.thy - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich +(* Title: HOL/TLA/Memory/MemoryImplementation.thy + Author: Stephan Merz, University of Munich *) header {* RPC-Memory example: Memory implementation *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/TLA/Memory/MemoryParameters.thy --- a/src/HOL/TLA/Memory/MemoryParameters.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/TLA/Memory/MemoryParameters.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,8 +1,5 @@ -(* - File: MemoryParameters.thy - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich +(* Title: HOL/TLA/Memory/MemoryParameters.thy + Author: Stephan Merz, University of Munich *) header {* RPC-Memory example: Memory parameters *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/TLA/Memory/ProcedureInterface.thy --- a/src/HOL/TLA/Memory/ProcedureInterface.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/TLA/Memory/ProcedureInterface.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,8 +1,5 @@ -(* - File: ProcedureInterface.thy - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich +(* Title: HOL/TLA/Memory/ProcedureInterface.thy + Author: Stephan Merz, University of Munich *) header {* Procedure interface for RPC-Memory components *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/TLA/Memory/RPC.thy --- a/src/HOL/TLA/Memory/RPC.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/TLA/Memory/RPC.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,8 +1,5 @@ -(* - File: RPC.thy - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich +(* Title: HOL/TLA/Memory/RPC.thy + Author: Stephan Merz, University of Munich *) header {* RPC-Memory example: RPC specification *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/TLA/Memory/RPCMemoryParams.thy --- a/src/HOL/TLA/Memory/RPCMemoryParams.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/TLA/Memory/RPCMemoryParams.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,8 +1,5 @@ -(* - File: RPCMemoryParams.thy - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich +(* Title: HOL/TLA/Memory/RPCMemoryParams.thy + Author: Stephan Merz, University of Munich *) header {* Basic declarations for the RPC-memory example *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/TLA/Memory/RPCParameters.thy --- a/src/HOL/TLA/Memory/RPCParameters.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/TLA/Memory/RPCParameters.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,8 +1,5 @@ -(* - File: RPCParameters.thy - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich +(* Title: HOL/TLA/Memory/RPCParameters.thy + Author: Stephan Merz, University of Munich *) header {* RPC-Memory example: RPC parameters *} diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Tools/SMT/smt_setup_solvers.ML --- a/src/HOL/Tools/SMT/smt_setup_solvers.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Tools/SMT/smt_setup_solvers.ML Tue Jan 25 09:45:45 2011 +0100 @@ -93,6 +93,15 @@ (* Z3 *) local + val flagN = "Z3_NON_COMMERCIAL" + + fun z3_make_command is_remote name () = + if getenv flagN = "yes" then make_command is_remote name () + else + error ("The SMT solver Z3 is not enabled. To enable it, set " ^ + "the environment variable " ^ quote flagN ^ " to " ^ quote ("yes") ^ + ".") + fun z3_options ctxt = ["-rs:" ^ string_of_int (Config.get ctxt SMT_Config.random_seed), "MODEL=true", "-smt"] @@ -117,7 +126,7 @@ name = make_name is_remote "z3", class = Z3_Interface.smtlib_z3C, avail = make_avail is_remote "Z3", - command = make_command is_remote "Z3", + command = z3_make_command is_remote "Z3", options = z3_options, default_max_relevant = 225, supports_filter = true, diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Tools/SMT/smt_solver.ML --- a/src/HOL/Tools/SMT/smt_solver.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Tools/SMT/smt_solver.ML Tue Jan 25 09:45:45 2011 +0100 @@ -55,7 +55,7 @@ local fun make_cmd command options problem_path proof_path = space_implode " " ( - map File.shell_quote (command @ options) @ + map File.shell_quote (command () @ options) @ [File.shell_path problem_path, "2>&1", ">", File.shell_path proof_path]) fun trace_and ctxt msg f x = @@ -136,7 +136,7 @@ |> tap (trace_assms ctxt) |> SMT_Translate.translate ctxt comments ||> tap trace_recon_data - in (run_solver ctxt' name (make_cmd (command ()) options) str, recon) end + in (run_solver ctxt' name (make_cmd command options) str, recon) end end diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Tools/SMT/z3_model.ML --- a/src/HOL/Tools/SMT/z3_model.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Tools/SMT/z3_model.ML Tue Jan 25 09:45:45 2011 +0100 @@ -196,15 +196,16 @@ let fun part _ [] = NONE | part us (t :: ts) = - (case try HOLogic.dest_eq t of - SOME lr => - if pred lr then SOME (lr, fold cons us ts) else part (t :: us) ts - | NONE => part (t :: us) ts) + (case try (pred o HOLogic.dest_eq) t of + SOME (SOME lr) => SOME (lr, fold cons us ts) + | _ => part (t :: us) ts) in (fn ts => part [] ts) end fun replace_vars tab = let - fun replace (v as Var _) = the_default v (AList.lookup (op aconv) tab v) + fun repl v = the_default v (AList.lookup (op aconv) tab v) + fun replace (v as Var _) = repl v + | replace (v as Free _) = repl v | replace t = t in map (Term.map_aterms replace) end @@ -241,21 +242,38 @@ in (map unfold_eq eqs, filter_out is_fun_app defs) end -fun unfold_eqs (eqs, defs) = +val unfold_eqs = let val is_ground = not o Term.exists_subterm Term.is_Var + fun is_non_rec (v, t) = not (Term.exists_subterm (equal v) t) + + fun rewr_var (l as Var _, r) = if is_ground r then SOME (l, r) else NONE + | rewr_var (r, l as Var _) = if is_ground r then SOME (l, r) else NONE + | rewr_var _ = NONE + + fun rewr_free' e = if is_non_rec e then SOME e else NONE + fun rewr_free (e as (Free _, _)) = rewr_free' e + | rewr_free (e as (_, Free _)) = rewr_free' (swap e) + | rewr_free _ = NONE fun is_trivial (Const (@{const_name HOL.eq}, _) $ t $ u) = t aconv u | is_trivial _ = false fun replace r = replace_vars [r] #> filter_out is_trivial - fun unfold (es, ds) = - (case first_eq (fn (l, Var _) => is_ground l | _ => false) es of - SOME ((l, r), es') => unfold (pairself (replace (r, l)) (es', ds)) + fun unfold_vars (es, ds) = + (case first_eq rewr_var es of + SOME (lr, es') => unfold_vars (pairself (replace lr) (es', ds)) | NONE => (es, ds)) - in unfold (eqs, defs) end + fun unfold_frees ues (es, ds) = + (case first_eq rewr_free es of + SOME (lr, es') => + pairself (replace lr) (es', ds) + |> unfold_frees (HOLogic.mk_eq lr :: replace lr ues) + | NONE => (ues @ es, ds)) + + in unfold_vars #> unfold_frees [] end fun swap_free ((eq as Const (@{const_name HOL.eq}, _)) $ t $ (u as Free _)) = eq $ u $ t diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Tools/TFL/dcterm.ML --- a/src/HOL/Tools/TFL/dcterm.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Tools/TFL/dcterm.ML Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/Tools/TFL/dcterm.ML - ID: $Id$ Author: Konrad Slind, Cambridge University Computer Laboratory - Copyright 1997 University of Cambridge *) (*--------------------------------------------------------------------------- diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Tools/TFL/thry.ML --- a/src/HOL/Tools/TFL/thry.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Tools/TFL/thry.ML Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/Tools/TFL/thry.ML - ID: $Id$ Author: Konrad Slind, Cambridge University Computer Laboratory - Copyright 1997 University of Cambridge *) signature THRY = diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Tools/TFL/utils.ML --- a/src/HOL/Tools/TFL/utils.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Tools/TFL/utils.ML Tue Jan 25 09:45:45 2011 +0100 @@ -1,7 +1,5 @@ (* Title: HOL/Tools/TFL/utils.ML - ID: $Id$ Author: Konrad Slind, Cambridge University Computer Laboratory - Copyright 1997 University of Cambridge Basic utilities. *) diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Tools/record.ML --- a/src/HOL/Tools/record.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Tools/record.ML Tue Jan 25 09:45:45 2011 +0100 @@ -25,6 +25,7 @@ cases: thm, simps: thm list, iffs: thm list} val get_info: theory -> string -> info option val the_info: theory -> string -> info + val get_hierarchy: theory -> (string * typ list) -> (string * ((string * sort) * typ) list) list val add_record: bool -> (string * sort) list * binding -> (typ list * string) option -> (binding * typ * mixfix) list -> theory -> theory @@ -93,8 +94,8 @@ fun merge data = Symtab.merge Thm.eq_thm_prop data; (* FIXME handle Symtab.DUP ?? *) ); -fun get_typedef_info tyco vs (({ rep_type, Abs_name, Rep_name, ...}, - { Abs_inject, Rep_inject, Abs_inverse, Rep_inverse, ... }) : Typedef.info) thy = +fun get_typedef_info tyco vs + (({rep_type, Abs_name, ...}, {Rep_inject, Abs_inverse, ... }) : Typedef.info) thy = let val exists_thm = UNIV_I @@ -196,8 +197,6 @@ struct val eq_reflection = @{thm eq_reflection}; -val atomize_all = @{thm HOL.atomize_all}; -val atomize_imp = @{thm HOL.atomize_imp}; val meta_allE = @{thm Pure.meta_allE}; val prop_subst = @{thm prop_subst}; val K_record_comp = @{thm K_record_comp}; @@ -616,12 +615,14 @@ (* parent records *) -fun add_parents _ NONE parents = parents - | add_parents thy (SOME (types, name)) parents = +local + +fun add_parents _ NONE = I + | add_parents thy (SOME (types, name)) = let fun err msg = error (msg ^ " parent record " ^ quote name); - val {args, parent, fields, extension, induct_scheme, ext_def, ...} = + val {args, parent, ...} = (case get_info thy name of SOME info => info | NONE => err "Unknown"); val _ = if length types <> length args then err "Bad number of arguments for" else (); @@ -630,15 +631,25 @@ val bads = map_filter bad_inst (args ~~ types); val _ = null bads orelse err ("Ill-sorted instantiation of " ^ commas bads ^ " in"); - val inst = map fst args ~~ types; - val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst); + val inst = args ~~ types; + val subst = Term.map_type_tfree (the o AList.lookup (op =) inst); val parent' = Option.map (apfst (map subst)) parent; - val fields' = map (apsnd subst) fields; - val extension' = apsnd (map subst) extension; - in - add_parents thy parent' - (make_parent_info name fields' extension' ext_def induct_scheme :: parents) - end; + in cons (name, inst) #> add_parents thy parent' end; + +in + +fun get_hierarchy thy (name, types) = add_parents thy (SOME (types, name)) []; + +fun get_parent_info thy parent = + add_parents thy parent [] |> map (fn (name, inst) => + let + val subst = Term.map_type_tfree (the o AList.lookup (op =) inst); + val {fields, extension, induct_scheme, ext_def, ...} = the_info thy name; + val fields' = map (apsnd subst) fields; + val extension' = apsnd (map subst) extension; + in make_parent_info name fields' extension' ext_def induct_scheme end); + +end; @@ -942,26 +953,30 @@ local +fun dest_update ctxt c = + (case try Syntax.unmark_const c of + SOME d => try (unsuffix updateN) (Consts.extern (ProofContext.consts_of ctxt) d) + | NONE => NONE); + fun field_updates_tr' ctxt (tm as Const (c, _) $ k $ u) = - let - val extern = Consts.extern (ProofContext.consts_of ctxt); - val t = - (case k of - Abs (_, _, Abs (_, _, t) $ Bound 0) => - if null (loose_bnos t) then t else raise Match - | Abs (_, _, t) => - if null (loose_bnos t) then t else raise Match - | _ => raise Match); - in - (case Option.map extern (try Syntax.unmark_const c) of - SOME update_name => - (case try (unsuffix updateN) update_name of - SOME name => + (case dest_update ctxt c of + SOME name => + let + val opt_t = + (case k of + Abs (_, _, Abs (_, _, t) $ Bound 0) => + if null (loose_bnos t) then SOME t else NONE + | Abs (_, _, t) => + if null (loose_bnos t) then SOME t else NONE + | _ => NONE); + in + (case opt_t of + SOME t => apfst (cons (Syntax.const @{syntax_const "_field_update"} $ Syntax.free name $ t)) (field_updates_tr' ctxt u) | NONE => ([], tm)) - | NONE => ([], tm)) - end + end + | NONE => ([], tm)) | field_updates_tr' _ tm = ([], tm); fun record_update_tr' ctxt tm = @@ -1509,11 +1524,6 @@ (* prepare arguments *) -fun read_raw_parent ctxt raw_T = - (case ProofContext.read_typ_abbrev ctxt raw_T of - Type (name, Ts) => (Ts, name) - | T => error ("Bad parent record specification: " ^ Syntax.string_of_typ ctxt T)); - fun read_typ ctxt raw_T env = let val ctxt' = fold (Variable.declare_typ o TFree) env ctxt; @@ -1809,7 +1819,7 @@ ((HOLogic.mk_random T' size, @{typ Random.seed}), SOME (v, termifyT T'))) params) tc @{typ Random.seed} (SOME Tm, @{typ Random.seed}); val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); - in + in thy |> Class.instantiation ([tyco], vs, @{sort random}) |> `(fn lthy => Syntax.check_term lthy eq) @@ -1822,11 +1832,14 @@ let val algebra = Sign.classes_of thy; val has_inst = can (Sorts.mg_domain algebra ext_tyco) @{sort random}; - in if has_inst then thy - else case Quickcheck_Generators.perhaps_constrain thy (map (rpair @{sort random}) Ts) vs - of SOME constrain => instantiate_random_record ext_tyco (map constrain vs) extN - ((map o map_atyps) (fn TFree v => TFree (constrain v)) Ts) thy - | NONE => thy + in + if has_inst then thy + else + (case Quickcheck_Generators.perhaps_constrain thy (map (rpair @{sort random}) Ts) vs of + SOME constrain => + instantiate_random_record ext_tyco (map constrain vs) extN + ((map o map_atyps) (fn TFree v => TFree (constrain v)) Ts) thy + | NONE => thy) end; fun add_code ext_tyco vs extT ext simps inject thy = @@ -1844,7 +1857,7 @@ fun mk_eq_refl thy = @{thm equal_refl} |> Thm.instantiate - ([pairself (Thm.ctyp_of thy) (TVar (("'a", 0), @{sort equal}), Logic.varifyT_global extT)], []) + ([pairself (ctyp_of thy) (TVar (("'a", 0), @{sort equal}), Logic.varifyT_global extT)], []) |> AxClass.unoverload thy; in thy @@ -1853,11 +1866,12 @@ |> Class.instantiation ([ext_tyco], vs, [HOLogic.class_equal]) |> `(fn lthy => Syntax.check_term lthy eq) |-> (fn eq => Specification.definition - (NONE, (Attrib.empty_binding, eq))) + (NONE, (Attrib.empty_binding, eq))) |-> (fn (_, (_, eq_def)) => Class.prove_instantiation_exit_result Morphism.thm - (fn _ => fn eq_def => tac eq_def) eq_def) - |-> (fn eq_def => fn thy => thy |> Code.del_eqn eq_def |> Code.add_default_eqn (mk_eq thy eq_def)) + (fn _ => fn eq_def => tac eq_def) eq_def) + |-> (fn eq_def => fn thy => + thy |> Code.del_eqn eq_def |> Code.add_default_eqn (mk_eq thy eq_def)) |> (fn thy => Code.add_nbe_default_eqn (mk_eq_refl thy) thy) |> ensure_random_record ext_tyco vs (fst ext) (binder_types (snd ext)) end; @@ -2418,7 +2432,7 @@ handle ERROR msg => cat_error msg ("The error(s) above occurred in parent record specification"); val parent_args = (case parent of SOME (Ts, _) => Ts | NONE => []); - val parents = add_parents thy parent []; + val parents = get_parent_info thy parent; val bfields = map (prep_field cert_typ) raw_fields; diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/Transcendental.thy Tue Jan 25 09:45:45 2011 +0100 @@ -164,7 +164,7 @@ { have "?s 0 = 0" by auto have Suc_m1: "\ n. Suc n - 1 = n" by auto - { fix B T E have "(if \ B then T else E) = (if B then E else T)" by auto } note if_eq = this + have if_eq: "\B T E. (if \ B then T else E) = (if B then E else T)" by auto have "?s sums y" using sums_if'[OF `f sums y`] . from this[unfolded sums_def, THEN LIMSEQ_Suc] @@ -348,7 +348,7 @@ fixes z :: "'a :: {monoid_mult,comm_ring}" shows "(\p=0..p=0..i = 0.. 'a set \ 'a set" (infixl "\" 65) where + "A \ B = {c. \a\A. \b\B. c = a + b}" + +definition set_times :: "'a::times set \ 'a set \ 'a set" (infixl "\" 70) where + "A \ B = {c. \a\A. \b\B. c = a * b}" + +definition elt_set_plus :: "'a::plus \ 'a set \ 'a set" (infixl "+o" 70) where + "a +o B = {c. \b\B. c = a + b}" + +definition elt_set_times :: "'a::times \ 'a set \ 'a set" (infixl "*o" 80) where + "a *o B = {c. \b\B. c = a * b}" + +abbreviation (input) elt_set_eq :: "'a \ 'a set \ bool" (infix "=o" 50) where + "x =o A \ x \ A" + +interpretation set_add!: semigroup "set_plus :: 'a::semigroup_add set \ 'a set \ 'a set" proof +qed (force simp add: set_plus_def add.assoc) + +interpretation set_add!: abel_semigroup "set_plus :: 'a::ab_semigroup_add set \ 'a set \ 'a set" proof +qed (force simp add: set_plus_def add.commute) + +interpretation set_add!: monoid "set_plus :: 'a::monoid_add set \ 'a set \ 'a set" "{0}" proof +qed (simp_all add: set_plus_def) + +interpretation set_add!: comm_monoid "set_plus :: 'a::comm_monoid_add set \ 'a set \ 'a set" "{0}" proof +qed (simp add: set_plus_def) + +interpretation set_add!: monoid_add "set_plus :: 'a::monoid_add set \ 'a set \ 'a set" "{0}" + defines listsum_set is set_add.listsum +proof +qed (simp_all add: set_add.assoc) + +interpretation set_add!: comm_monoid_add "set_plus :: 'a::comm_monoid_add set \ 'a set \ 'a set" "{0}" + defines setsum_set is set_add.setsum + where "monoid_add.listsum set_plus {0::'a} = listsum_set" +proof - + show "class.comm_monoid_add (set_plus :: 'a set \ 'a set \ 'a set) {0}" proof + qed (simp_all add: set_add.commute) + then interpret set_add!: comm_monoid_add "set_plus :: 'a set \ 'a set \ 'a set" "{0}" . + show "monoid_add.listsum set_plus {0::'a} = listsum_set" + by (simp only: listsum_set_def) +qed + +interpretation set_mult!: semigroup "set_times :: 'a::semigroup_mult set \ 'a set \ 'a set" proof +qed (force simp add: set_times_def mult.assoc) + +interpretation set_mult!: abel_semigroup "set_times :: 'a::ab_semigroup_mult set \ 'a set \ 'a set" proof +qed (force simp add: set_times_def mult.commute) + +interpretation set_mult!: monoid "set_times :: 'a::monoid_mult set \ 'a set \ 'a set" "{1}" proof +qed (simp_all add: set_times_def) + +interpretation set_mult!: comm_monoid "set_times :: 'a::comm_monoid_mult set \ 'a set \ 'a set" "{1}" proof +qed (simp add: set_times_def) + +interpretation set_mult!: monoid_mult "{1}" "set_times :: 'a::monoid_mult set \ 'a set \ 'a set" + defines power_set is set_mult.power +proof +qed (simp_all add: set_mult.assoc) + +interpretation set_mult!: comm_monoid_mult "set_times :: 'a::comm_monoid_mult set \ 'a set \ 'a set" "{1}" + defines setprod_set is set_mult.setprod + where "power.power {1} set_times = power_set" +proof - + show "class.comm_monoid_mult (set_times :: 'a set \ 'a set \ 'a set) {1}" proof + qed (simp_all add: set_mult.commute) + then interpret set_mult!: comm_monoid_mult "set_times :: 'a set \ 'a set \ 'a set" "{1}" . + show "power.power {1} set_times = power_set" + by (simp add: power_set_def) +qed + +lemma set_plus_intro [intro]: "a : C ==> b : D ==> a + b : C \ D" + by (auto simp add: set_plus_def) + +lemma set_plus_intro2 [intro]: "b : C ==> a + b : a +o C" + by (auto simp add: elt_set_plus_def) + +lemma set_plus_rearrange: "((a::'a::comm_monoid_add) +o C) \ + (b +o D) = (a + b) +o (C \ D)" + apply (auto simp add: elt_set_plus_def set_plus_def add_ac) + apply (rule_tac x = "ba + bb" in exI) + apply (auto simp add: add_ac) + apply (rule_tac x = "aa + a" in exI) + apply (auto simp add: add_ac) + done + +lemma set_plus_rearrange2: "(a::'a::semigroup_add) +o (b +o C) = + (a + b) +o C" + by (auto simp add: elt_set_plus_def add_assoc) + +lemma set_plus_rearrange3: "((a::'a::semigroup_add) +o B) \ C = + a +o (B \ C)" + apply (auto simp add: elt_set_plus_def set_plus_def) + apply (blast intro: add_ac) + apply (rule_tac x = "a + aa" in exI) + apply (rule conjI) + apply (rule_tac x = "aa" in bexI) + apply auto + apply (rule_tac x = "ba" in bexI) + apply (auto simp add: add_ac) + done + +theorem set_plus_rearrange4: "C \ ((a::'a::comm_monoid_add) +o D) = + a +o (C \ D)" + apply (auto intro!: subsetI simp add: elt_set_plus_def set_plus_def add_ac) + apply (rule_tac x = "aa + ba" in exI) + apply (auto simp add: add_ac) + done + +theorems set_plus_rearranges = set_plus_rearrange set_plus_rearrange2 + set_plus_rearrange3 set_plus_rearrange4 + +lemma set_plus_mono [intro!]: "C <= D ==> a +o C <= a +o D" + by (auto simp add: elt_set_plus_def) + +lemma set_plus_mono2 [intro]: "(C::('a::plus) set) <= D ==> E <= F ==> + C \ E <= D \ F" + by (auto simp add: set_plus_def) + +lemma set_plus_mono3 [intro]: "a : C ==> a +o D <= C \ D" + by (auto simp add: elt_set_plus_def set_plus_def) + +lemma set_plus_mono4 [intro]: "(a::'a::comm_monoid_add) : C ==> + a +o D <= D \ C" + by (auto simp add: elt_set_plus_def set_plus_def add_ac) + +lemma set_plus_mono5: "a:C ==> B <= D ==> a +o B <= C \ D" + apply (subgoal_tac "a +o B <= a +o D") + apply (erule order_trans) + apply (erule set_plus_mono3) + apply (erule set_plus_mono) + done + +lemma set_plus_mono_b: "C <= D ==> x : a +o C + ==> x : a +o D" + apply (frule set_plus_mono) + apply auto + done + +lemma set_plus_mono2_b: "C <= D ==> E <= F ==> x : C \ E ==> + x : D \ F" + apply (frule set_plus_mono2) + prefer 2 + apply force + apply assumption + done + +lemma set_plus_mono3_b: "a : C ==> x : a +o D ==> x : C \ D" + apply (frule set_plus_mono3) + apply auto + done + +lemma set_plus_mono4_b: "(a::'a::comm_monoid_add) : C ==> + x : a +o D ==> x : D \ C" + apply (frule set_plus_mono4) + apply auto + done + +lemma set_zero_plus [simp]: "(0::'a::comm_monoid_add) +o C = C" + by (auto simp add: elt_set_plus_def) + +lemma set_zero_plus2: "(0::'a::comm_monoid_add) : A ==> B <= A \ B" + apply (auto intro!: subsetI simp add: set_plus_def) + apply (rule_tac x = 0 in bexI) + apply (rule_tac x = x in bexI) + apply (auto simp add: add_ac) + done + +lemma set_plus_imp_minus: "(a::'a::ab_group_add) : b +o C ==> (a - b) : C" + by (auto simp add: elt_set_plus_def add_ac diff_minus) + +lemma set_minus_imp_plus: "(a::'a::ab_group_add) - b : C ==> a : b +o C" + apply (auto simp add: elt_set_plus_def add_ac diff_minus) + apply (subgoal_tac "a = (a + - b) + b") + apply (rule bexI, assumption, assumption) + apply (auto simp add: add_ac) + done + +lemma set_minus_plus: "((a::'a::ab_group_add) - b : C) = (a : b +o C)" + by (rule iffI, rule set_minus_imp_plus, assumption, rule set_plus_imp_minus, + assumption) + +lemma set_times_intro [intro]: "a : C ==> b : D ==> a * b : C \ D" + by (auto simp add: set_times_def) + +lemma set_times_intro2 [intro!]: "b : C ==> a * b : a *o C" + by (auto simp add: elt_set_times_def) + +lemma set_times_rearrange: "((a::'a::comm_monoid_mult) *o C) \ + (b *o D) = (a * b) *o (C \ D)" + apply (auto simp add: elt_set_times_def set_times_def) + apply (rule_tac x = "ba * bb" in exI) + apply (auto simp add: mult_ac) + apply (rule_tac x = "aa * a" in exI) + apply (auto simp add: mult_ac) + done + +lemma set_times_rearrange2: "(a::'a::semigroup_mult) *o (b *o C) = + (a * b) *o C" + by (auto simp add: elt_set_times_def mult_assoc) + +lemma set_times_rearrange3: "((a::'a::semigroup_mult) *o B) \ C = + a *o (B \ C)" + apply (auto simp add: elt_set_times_def set_times_def) + apply (blast intro: mult_ac) + apply (rule_tac x = "a * aa" in exI) + apply (rule conjI) + apply (rule_tac x = "aa" in bexI) + apply auto + apply (rule_tac x = "ba" in bexI) + apply (auto simp add: mult_ac) + done + +theorem set_times_rearrange4: "C \ ((a::'a::comm_monoid_mult) *o D) = + a *o (C \ D)" + apply (auto intro!: subsetI simp add: elt_set_times_def set_times_def + mult_ac) + apply (rule_tac x = "aa * ba" in exI) + apply (auto simp add: mult_ac) + done + +theorems set_times_rearranges = set_times_rearrange set_times_rearrange2 + set_times_rearrange3 set_times_rearrange4 + +lemma set_times_mono [intro]: "C <= D ==> a *o C <= a *o D" + by (auto simp add: elt_set_times_def) + +lemma set_times_mono2 [intro]: "(C::('a::times) set) <= D ==> E <= F ==> + C \ E <= D \ F" + by (auto simp add: set_times_def) + +lemma set_times_mono3 [intro]: "a : C ==> a *o D <= C \ D" + by (auto simp add: elt_set_times_def set_times_def) + +lemma set_times_mono4 [intro]: "(a::'a::comm_monoid_mult) : C ==> + a *o D <= D \ C" + by (auto simp add: elt_set_times_def set_times_def mult_ac) + +lemma set_times_mono5: "a:C ==> B <= D ==> a *o B <= C \ D" + apply (subgoal_tac "a *o B <= a *o D") + apply (erule order_trans) + apply (erule set_times_mono3) + apply (erule set_times_mono) + done + +lemma set_times_mono_b: "C <= D ==> x : a *o C + ==> x : a *o D" + apply (frule set_times_mono) + apply auto + done + +lemma set_times_mono2_b: "C <= D ==> E <= F ==> x : C \ E ==> + x : D \ F" + apply (frule set_times_mono2) + prefer 2 + apply force + apply assumption + done + +lemma set_times_mono3_b: "a : C ==> x : a *o D ==> x : C \ D" + apply (frule set_times_mono3) + apply auto + done + +lemma set_times_mono4_b: "(a::'a::comm_monoid_mult) : C ==> + x : a *o D ==> x : D \ C" + apply (frule set_times_mono4) + apply auto + done + +lemma set_one_times [simp]: "(1::'a::comm_monoid_mult) *o C = C" + by (auto simp add: elt_set_times_def) + +lemma set_times_plus_distrib: "(a::'a::semiring) *o (b +o C)= + (a * b) +o (a *o C)" + by (auto simp add: elt_set_plus_def elt_set_times_def ring_distribs) + +lemma set_times_plus_distrib2: "(a::'a::semiring) *o (B \ C) = + (a *o B) \ (a *o C)" + apply (auto simp add: set_plus_def elt_set_times_def ring_distribs) + apply blast + apply (rule_tac x = "b + bb" in exI) + apply (auto simp add: ring_distribs) + done + +lemma set_times_plus_distrib3: "((a::'a::semiring) +o C) \ D <= + a *o D \ C \ D" + apply (auto intro!: subsetI simp add: + elt_set_plus_def elt_set_times_def set_times_def + set_plus_def ring_distribs) + apply auto + done + +theorems set_times_plus_distribs = + set_times_plus_distrib + set_times_plus_distrib2 + +lemma set_neg_intro: "(a::'a::ring_1) : (- 1) *o C ==> + - a : C" + by (auto simp add: elt_set_times_def) + +lemma set_neg_intro2: "(a::'a::ring_1) : C ==> + - a : (- 1) *o C" + by (auto simp add: elt_set_times_def) + +lemma set_plus_image: + fixes S T :: "'n::semigroup_add set" shows "S \ T = (\(x, y). x + y) ` (S \ T)" + unfolding set_plus_def by (fastsimp simp: image_iff) + +lemma set_setsum_alt: + assumes fin: "finite I" + shows "setsum_set S I = {setsum s I |s. \i\I. s i \ S i}" + (is "_ = ?setsum I") +using fin proof induct + case (insert x F) + have "setsum_set S (insert x F) = S x \ ?setsum F" + using insert.hyps by auto + also have "...= {s x + setsum s F |s. \ i\insert x F. s i \ S i}" + unfolding set_plus_def + proof safe + fix y s assume "y \ S x" "\i\F. s i \ S i" + then show "\s'. y + setsum s F = s' x + setsum s' F \ (\i\insert x F. s' i \ S i)" + using insert.hyps + by (intro exI[of _ "\i. if i \ F then s i else y"]) (auto simp add: set_plus_def) + qed auto + finally show ?case + using insert.hyps by auto +qed auto + +lemma setsum_set_cond_linear: + fixes f :: "('a::comm_monoid_add) set \ ('b::comm_monoid_add) set" + assumes [intro!]: "\A B. P A \ P B \ P (A \ B)" "P {0}" + and f: "\A B. P A \ P B \ f (A \ B) = f A \ f B" "f {0} = {0}" + assumes all: "\i. i \ I \ P (S i)" + shows "f (setsum_set S I) = setsum_set (f \ S) I" +proof cases + assume "finite I" from this all show ?thesis + proof induct + case (insert x F) + from `finite F` `\i. i \ insert x F \ P (S i)` have "P (setsum_set S F)" + by induct auto + with insert show ?case + by (simp, subst f) auto + qed (auto intro!: f) +qed (auto intro!: f) + +lemma setsum_set_linear: + fixes f :: "('a::comm_monoid_add) set => ('b::comm_monoid_add) set" + assumes "\A B. f(A) \ f(B) = f(A \ B)" "f {0} = {0}" + shows "f (setsum_set S I) = setsum_set (f \ S) I" + using setsum_set_cond_linear[of "\x. True" f I S] assms by auto + +end diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/ex/Sudoku.thy --- a/src/HOL/ex/Sudoku.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/ex/Sudoku.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,5 +1,4 @@ -(* Title: Sudoku.thy - ID: $Id$ +(* Title: HOL/ex/Sudoku.thy Author: Tjark Weber Copyright 2005-2008 *) diff -r baf1964bc468 -r 621fa31e1dbd src/HOL/ex/svc_test.thy --- a/src/HOL/ex/svc_test.thy Mon Jan 24 22:29:50 2011 +0100 +++ b/src/HOL/ex/svc_test.thy Tue Jan 25 09:45:45 2011 +0100 @@ -1,6 +1,3 @@ - -(* $Id$ *) - header {* Demonstrating the interface SVC *} theory svc_test diff -r baf1964bc468 -r 621fa31e1dbd src/Provers/classical.ML --- a/src/Provers/classical.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/Provers/classical.ML Tue Jan 25 09:45:45 2011 +0100 @@ -153,7 +153,7 @@ *) fun classical_rule rule = - if Object_Logic.is_elim rule then + if is_some (Object_Logic.elim_concl rule) then let val rule' = rule RS classical; val concl' = Thm.concl_of rule'; diff -r baf1964bc468 -r 621fa31e1dbd src/Pure/Isar/class_declaration.ML --- a/src/Pure/Isar/class_declaration.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/Pure/Isar/class_declaration.ML Tue Jan 25 09:45:45 2011 +0100 @@ -6,13 +6,14 @@ signature CLASS_DECLARATION = sig - val class: binding -> class list -> Element.context_i list - -> theory -> string * local_theory - val class_cmd: binding -> xstring list -> Element.context list - -> theory -> string * local_theory - val prove_subclass: tactic -> class -> local_theory -> local_theory - val subclass: class -> local_theory -> Proof.state - val subclass_cmd: xstring -> local_theory -> Proof.state + val class: (local_theory -> local_theory) -> binding -> class list -> + Element.context_i list -> theory -> string * local_theory + val class_cmd: (local_theory -> local_theory) -> binding -> xstring list -> + Element.context list -> theory -> string * local_theory + val prove_subclass: (local_theory -> local_theory) -> tactic -> class -> + local_theory -> local_theory + val subclass: (local_theory -> local_theory) -> class -> local_theory -> Proof.state + val subclass_cmd: (local_theory -> local_theory) -> xstring -> local_theory -> Proof.state end; structure Class_Declaration: CLASS_DECLARATION = @@ -288,14 +289,14 @@ #> pair (param_map, params, assm_axiom))) end; -fun gen_class prep_class_spec b raw_supclasses raw_elems thy = +fun gen_class prep_class_spec before_exit b raw_supclasses raw_elems thy = let val class = Sign.full_name thy b; val (((sups, supparam_names), (supsort, base_sort, supexpr)), (elems, global_syntax)) = prep_class_spec thy raw_supclasses raw_elems; in thy - |> Expression.add_locale b (Binding.qualify true "class" b) supexpr elems + |> Expression.add_locale I b (Binding.qualify true "class" b) supexpr elems |> snd |> Local_Theory.exit_global |> adjungate_axclass b class base_sort sups supsort supparam_names global_syntax ||> Theory.checkpoint @@ -305,7 +306,7 @@ Context.theory_map (Locale.add_registration (class, base_morph) (Option.map (rpair true) eq_morph) export_morph) #> Class.register class sups params base_sort base_morph export_morph axiom assm_intro of_class)) - |> Named_Target.init class + |> Named_Target.init before_exit class |> pair class end; @@ -321,7 +322,7 @@ local -fun gen_subclass prep_class do_proof raw_sup lthy = +fun gen_subclass prep_class do_proof before_exit raw_sup lthy = let val thy = ProofContext.theory_of lthy; val proto_sup = prep_class thy raw_sup; @@ -338,7 +339,7 @@ fun after_qed some_wit = ProofContext.background_theory (Class.register_subclass (sub, sup) some_dep_morph some_wit export) - #> ProofContext.theory_of #> Named_Target.init sub; + #> ProofContext.theory_of #> Named_Target.init before_exit sub; in do_proof after_qed some_prop goal_ctxt end; fun user_proof after_qed some_prop = @@ -352,7 +353,7 @@ in val subclass = gen_subclass (K I) user_proof; -fun prove_subclass tac = gen_subclass (K I) (tactic_proof tac); +fun prove_subclass before_exit tac = gen_subclass (K I) (tactic_proof tac) before_exit; val subclass_cmd = gen_subclass (ProofContext.read_class o ProofContext.init_global) user_proof; end; (*local*) diff -r baf1964bc468 -r 621fa31e1dbd src/Pure/Isar/element.ML --- a/src/Pure/Isar/element.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/Pure/Isar/element.ML Tue Jan 25 09:45:45 2011 +0100 @@ -201,6 +201,16 @@ local +fun standard_elim th = + (case Object_Logic.elim_concl th of + SOME C => + let + val cert = Thm.cterm_of (Thm.theory_of_thm th); + val thesis = Var ((Auto_Bind.thesisN, Thm.maxidx_of th + 1), fastype_of C); + val th' = Thm.instantiate ([], [(cert C, cert thesis)]) th; + in (th', true) end + | NONE => (th, false)); + fun thm_name kind th prts = let val head = if Thm.has_name_hint th then @@ -209,13 +219,13 @@ else Pretty.command kind in Pretty.block (Pretty.fbreaks (head :: prts)) end; -fun fix (x, T) = (Binding.name x, SOME T); - fun obtain prop ctxt = let - val ((xs, prop'), ctxt') = Variable.focus prop ctxt; + val ((ps, prop'), ctxt') = Variable.focus prop ctxt; + fun fix (x, T) = (Binding.name (ProofContext.revert_skolem ctxt' x), SOME T); + val xs = map (fix o Term.dest_Free o Thm.term_of o #2) ps; val As = Logic.strip_imp_prems (Thm.term_of prop'); - in ((Binding.empty, (map (fix o Term.dest_Free o Thm.term_of o #2) xs, As)), ctxt') end; + in ((Binding.empty, (xs, As)), ctxt') end; in @@ -224,17 +234,15 @@ val thy = ProofContext.theory_of ctxt; val cert = Thm.cterm_of thy; - val th = Raw_Simplifier.norm_hhf raw_th; - val is_elim = Object_Logic.is_elim th; - - val ((_, [th']), ctxt') = Variable.import true [th] (Variable.set_body false ctxt); + val (th, is_elim) = standard_elim (Raw_Simplifier.norm_hhf raw_th); + val ((_, [th']), ctxt') = Variable.import true [th] (Variable.set_body true ctxt); val prop = Thm.prop_of th'; val (prems, concl) = Logic.strip_horn prop; val concl_term = Object_Logic.drop_judgment thy concl; val fixes = fold_aterms (fn v as Free (x, T) => if Variable.newly_fixed ctxt' ctxt x andalso not (v aconv concl_term) - then insert (op =) (x, T) else I | _ => I) prop [] |> rev; + then insert (op =) (ProofContext.revert_skolem ctxt' x, T) else I | _ => I) prop [] |> rev; val (assumes, cases) = take_suffix (fn prem => is_elim andalso concl aconv Logic.strip_assums_concl prem) prems; in diff -r baf1964bc468 -r 621fa31e1dbd src/Pure/Isar/expression.ML --- a/src/Pure/Isar/expression.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/Pure/Isar/expression.ML Tue Jan 25 09:45:45 2011 +0100 @@ -29,20 +29,20 @@ val read_declaration: expression -> (Proof.context -> Proof.context) -> Element.context list -> Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list * Element.context_i list) * ((string * typ) list * Proof.context) - val add_locale: binding -> binding -> expression_i -> Element.context_i list -> - theory -> string * local_theory - val add_locale_cmd: binding -> binding -> expression -> Element.context list -> - theory -> string * local_theory + val add_locale: (local_theory -> local_theory) -> binding -> binding -> + expression_i -> Element.context_i list -> theory -> string * local_theory + val add_locale_cmd: (local_theory -> local_theory) -> binding -> binding -> + expression -> Element.context list -> theory -> string * local_theory (* Interpretation *) val cert_goal_expression: expression_i -> Proof.context -> (term list list * (string * morphism) list * morphism) * Proof.context val read_goal_expression: expression -> Proof.context -> (term list list * (string * morphism) list * morphism) * Proof.context - val sublocale: string -> expression_i -> (Attrib.binding * term) list -> - theory -> Proof.state - val sublocale_cmd: string -> expression -> (Attrib.binding * string) list -> - theory -> Proof.state + val sublocale: (local_theory -> local_theory) -> string -> expression_i -> + (Attrib.binding * term) list -> theory -> Proof.state + val sublocale_cmd: (local_theory -> local_theory) -> string -> expression -> + (Attrib.binding * string) list -> theory -> Proof.state val interpretation: expression_i -> (Attrib.binding * term) list -> theory -> Proof.state val interpretation_cmd: expression -> (Attrib.binding * string) list -> @@ -731,7 +731,7 @@ | defines_to_notes _ e = e; fun gen_add_locale prep_decl - binding raw_predicate_binding raw_import raw_body thy = + before_exit binding raw_predicate_binding raw_import raw_body thy = let val name = Sign.full_name thy binding; val _ = Locale.defined thy name andalso @@ -784,7 +784,7 @@ val loc_ctxt = thy' |> Locale.register_locale binding (extraTs, params) (asm, rev defs) (a_intro, b_intro) axioms [] (rev notes) (rev deps') - |> Named_Target.init name + |> Named_Target.init before_exit name |> fold (fn (kind, facts) => Local_Theory.notes_kind kind facts #> snd) notes'; in (name, loc_ctxt) end; @@ -900,11 +900,11 @@ export theory) (deps ~~ witss)) end; -fun gen_sublocale prep_expr intern parse_prop prep_attr raw_target - expression equations thy = +fun gen_sublocale prep_expr intern parse_prop prep_attr + before_exit raw_target expression equations thy = let val target = intern thy raw_target; - val target_ctxt = Named_Target.init target thy; + val target_ctxt = Named_Target.init before_exit target thy; val ((propss, deps, export), expr_ctxt) = prep_expr expression target_ctxt; val eqns = map (parse_prop expr_ctxt o snd) equations |> Syntax.check_terms expr_ctxt; @@ -919,8 +919,8 @@ in fun sublocale x = gen_sublocale cert_goal_expression (K I) (K I) (K I) x; -fun sublocale_cmd x = gen_sublocale read_goal_expression Locale.intern - Syntax.parse_prop Attrib.intern_src x; +fun sublocale_cmd x = + gen_sublocale read_goal_expression Locale.intern Syntax.parse_prop Attrib.intern_src x; end; diff -r baf1964bc468 -r 621fa31e1dbd src/Pure/Isar/isar_syn.ML --- a/src/Pure/Isar/isar_syn.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/Pure/Isar/isar_syn.ML Tue Jan 25 09:45:45 2011 +0100 @@ -413,7 +413,7 @@ Scan.optional (Parse.$$$ "=" |-- Parse.!!! locale_val) (([], []), []) -- Parse.opt_begin >> (fn ((name, (expr, elems)), begin) => (begin ? Toplevel.print) o Toplevel.begin_local_theory begin - (Expression.add_locale_cmd name Binding.empty expr elems #> snd))); + (Expression.add_locale_cmd I name Binding.empty expr elems #> snd))); fun parse_interpretation_arguments mandatory = Parse.!!! (Parse_Spec.locale_expression mandatory) -- @@ -426,7 +426,7 @@ (Parse.xname --| (Parse.$$$ "\\" || Parse.$$$ "<") -- parse_interpretation_arguments false >> (fn (loc, (expr, equations)) => - Toplevel.print o Toplevel.theory_to_proof (Expression.sublocale_cmd loc expr equations))); + Toplevel.print o Toplevel.theory_to_proof (Expression.sublocale_cmd I loc expr equations))); val _ = Outer_Syntax.command "interpretation" @@ -456,11 +456,11 @@ (Parse.binding -- Scan.optional (Parse.$$$ "=" |-- class_val) ([], []) -- Parse.opt_begin >> (fn ((name, (supclasses, elems)), begin) => (begin ? Toplevel.print) o Toplevel.begin_local_theory begin - (Class_Declaration.class_cmd name supclasses elems #> snd))); + (Class_Declaration.class_cmd I name supclasses elems #> snd))); val _ = Outer_Syntax.local_theory_to_proof "subclass" "prove a subclass relation" Keyword.thy_goal - (Parse.xname >> Class_Declaration.subclass_cmd); + (Parse.xname >> Class_Declaration.subclass_cmd I); val _ = Outer_Syntax.command "instantiation" "instantiate and prove type arity" Keyword.thy_decl diff -r baf1964bc468 -r 621fa31e1dbd src/Pure/Isar/named_target.ML --- a/src/Pure/Isar/named_target.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/Pure/Isar/named_target.ML Tue Jan 25 09:45:45 2011 +0100 @@ -7,7 +7,7 @@ signature NAMED_TARGET = sig - val init: string -> theory -> local_theory + val init: (local_theory -> local_theory) -> string -> theory -> local_theory val theory_init: theory -> local_theory val reinit: local_theory -> local_theory -> local_theory val context_cmd: xstring -> theory -> local_theory @@ -19,12 +19,18 @@ (* context data *) -datatype target = Target of {target: string, is_locale: bool, is_class: bool}; +datatype target = + Target of {target: string, is_locale: bool, is_class: bool, + before_exit: local_theory -> local_theory}; -fun named_target _ "" = Target {target = "", is_locale = false, is_class = false} - | named_target thy locale = +fun make_target target is_locale is_class before_exit = + Target {target = target, is_locale = is_locale, is_class = is_class, + before_exit = before_exit}; + +fun named_target _ "" before_exit = make_target "" false false before_exit + | named_target thy locale before_exit = if Locale.defined thy locale - then Target {target = locale, is_locale = true, is_class = Class.is_class thy locale} + then make_target locale true (Class.is_class thy locale) before_exit else error ("No such locale: " ^ quote locale); structure Data = Proof_Data @@ -33,7 +39,9 @@ fun init _ = NONE; ); -val peek = Option.map (fn Target args => args) o Data.get; +val peek = + Data.get #> Option.map (fn Target {target, is_locale, is_class, ...} => + {target = target, is_locale = is_locale, is_class = is_class}); (* generic declarations *) @@ -169,14 +177,14 @@ (* init *) -fun init_context (Target {target, is_locale, is_class}) = +fun init_context (Target {target, is_locale, is_class, ...}) = if not is_locale then ProofContext.init_global else if not is_class then Locale.init target else Class.init target; -fun init target thy = +fun init before_exit target thy = let - val ta = named_target thy target; + val ta = named_target thy target before_exit; in thy |> init_context ta @@ -190,17 +198,17 @@ syntax_declaration = fn pervasive => target_declaration ta {syntax = true, pervasive = pervasive}, pretty = pretty ta, - exit = Local_Theory.target_of} + exit = Local_Theory.target_of o before_exit} end; -val theory_init = init ""; +val theory_init = init I ""; fun reinit lthy = - (case peek lthy of - SOME {target, ...} => init target o Local_Theory.exit_global + (case Data.get lthy of + SOME (Target {target, before_exit, ...}) => init before_exit target o Local_Theory.exit_global | NONE => error "Not in a named target"); -fun context_cmd "-" thy = init "" thy - | context_cmd target thy = init (Locale.intern thy target) thy; +fun context_cmd "-" thy = init I "" thy + | context_cmd target thy = init I (Locale.intern thy target) thy; end; diff -r baf1964bc468 -r 621fa31e1dbd src/Pure/Isar/object_logic.ML --- a/src/Pure/Isar/object_logic.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/Pure/Isar/object_logic.ML Tue Jan 25 09:45:45 2011 +0100 @@ -17,7 +17,7 @@ val ensure_propT: theory -> term -> term val dest_judgment: cterm -> cterm val judgment_conv: conv -> conv - val is_elim: thm -> bool + val elim_concl: thm -> term option val declare_atomize: attribute val declare_rulify: attribute val atomize_term: theory -> term -> term @@ -145,13 +145,15 @@ (* elimination rules *) -fun is_elim rule = +fun elim_concl rule = let val thy = Thm.theory_of_thm rule; val concl = Thm.concl_of rule; + val C = drop_judgment thy concl; in - Term.is_Var (drop_judgment thy concl) andalso + if Term.is_Var C andalso exists (fn prem => concl aconv Logic.strip_assums_concl prem) (Thm.prems_of rule) + then SOME C else NONE end; diff -r baf1964bc468 -r 621fa31e1dbd src/Pure/Isar/proof_display.ML --- a/src/Pure/Isar/proof_display.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/Pure/Isar/proof_display.ML Tue Jan 25 09:45:45 2011 +0100 @@ -80,12 +80,13 @@ local -fun pretty_fact_name (kind, "") = Pretty.str kind - | pretty_fact_name (kind, name) = Pretty.block [Pretty.str kind, Pretty.brk 1, - Pretty.str (Long_Name.base_name name), Pretty.str ":"]; +fun pretty_fact_name (kind, "") = Pretty.command kind + | pretty_fact_name (kind, name) = + Pretty.block [Pretty.command kind, Pretty.brk 1, + Pretty.str (Long_Name.base_name name), Pretty.str ":"]; fun pretty_facts ctxt = - flat o (separate [Pretty.fbrk, Pretty.str "and "]) o + flat o (separate [Pretty.fbrk, Pretty.keyword "and", Pretty.str " "]) o map (single o ProofContext.pretty_fact_aux ctxt false); in diff -r baf1964bc468 -r 621fa31e1dbd src/Pure/Thy/thm_deps.ML --- a/src/Pure/Thy/thm_deps.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/Pure/Thy/thm_deps.ML Tue Jan 25 09:45:45 2011 +0100 @@ -65,15 +65,14 @@ val used = Proofterm.fold_body_thms - (fn (a, prop, _) => a <> "" ? Symtab.insert_list (op =) (a, prop)) + (fn (a, _, _) => a <> "" ? Symtab.update (a, ())) (map (Proofterm.strip_thm o Thm.proof_body_of o #1 o #2) new_thms) Symtab.empty; - fun is_unused (a, th) = - not (member (op aconv) (Symtab.lookup_list used a) (Thm.prop_of th)); + fun is_unused a = not (Symtab.defined used a); (* groups containing at least one used theorem *) - val used_groups = fold (fn (a, (th, _, group)) => - if is_unused (a, th) then I + val used_groups = fold (fn (a, (_, _, group)) => + if is_unused a then I else (case group of NONE => I @@ -82,7 +81,7 @@ val (thms', _) = fold (fn (a, (th, concealed, group)) => fn q as (thms, seen_groups) => if not concealed andalso member (op =) [Thm.theoremK, Thm.lemmaK, Thm.corollaryK] (Thm.get_kind th) andalso - is_unused (a, th) + is_unused a then (case group of NONE => ((a, th) :: thms, seen_groups) diff -r baf1964bc468 -r 621fa31e1dbd src/Pure/Thy/thy_info.ML --- a/src/Pure/Thy/thy_info.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/Pure/Thy/thy_info.ML Tue Jan 25 09:45:45 2011 +0100 @@ -225,12 +225,12 @@ val _ = kill_thy name; val _ = Output.urgent_message ("Loading theory " ^ quote name ^ required_by " " initiators); - val {master = (thy_path, _), ...} = deps; + val {master = (thy_path, _), imports} = deps; val dir = Path.dir thy_path; val pos = Path.position thy_path; val uses = map (apfst Path.explode) (#3 (Thy_Header.read pos text)); fun init _ = - Thy_Load.begin_theory dir name parent_thys uses + Thy_Load.begin_theory dir name imports parent_thys uses |> Present.begin_theory update_time dir uses; val (after_load, theory) = Outer_Syntax.load_thy name init pos text; @@ -324,7 +324,7 @@ val _ = kill_thy name; val _ = use_thys_wrt dir imports; val parent_thys = map (get_theory o base_name) imports; - in Thy_Load.begin_theory dir name parent_thys uses end; + in Thy_Load.begin_theory dir name imports parent_thys uses end; (* register theory *) @@ -334,7 +334,8 @@ val name = Context.theory_name theory; val master = Thy_Load.check_thy (Thy_Load.master_directory theory) name; val parents = map Context.theory_name (Theory.parents_of theory); - val deps = make_deps master parents; + val imports = Thy_Load.imports_of theory; + val deps = make_deps master imports; in NAMED_CRITICAL "Thy_Info" (fn () => (kill_thy name; diff -r baf1964bc468 -r 621fa31e1dbd src/Pure/Thy/thy_load.ML --- a/src/Pure/Thy/thy_load.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/Pure/Thy/thy_load.ML Tue Jan 25 09:45:45 2011 +0100 @@ -13,6 +13,7 @@ val set_master_path: Path.T -> unit val get_master_path: unit -> Path.T val master_directory: theory -> Path.T + val imports_of: theory -> string list val provide: Path.T * (Path.T * file_ident) -> theory -> theory val legacy_show_path: unit -> string list val legacy_add_path: string -> unit @@ -28,7 +29,7 @@ val provide_file: Path.T -> theory -> theory val use_ml: Path.T -> unit val exec_ml: Path.T -> generic_theory -> generic_theory - val begin_theory: Path.T -> string -> theory list -> (Path.T * bool) list -> theory + val begin_theory: Path.T -> string -> string list -> theory list -> (Path.T * bool) list -> theory end; structure Thy_Load: THY_LOAD = @@ -83,40 +84,42 @@ type files = {master_dir: Path.T, (*master directory of theory source*) + imports: string list, (*source specification of imports*) required: Path.T list, (*source path*) provided: (Path.T * (Path.T * file_ident)) list}; (*source path, physical path, identifier*) -fun make_files (master_dir, required, provided): files = - {master_dir = master_dir, required = required, provided = provided}; +fun make_files (master_dir, imports, required, provided): files = + {master_dir = master_dir, imports = imports, required = required, provided = provided}; structure Files = Theory_Data ( type T = files; - val empty = make_files (Path.current, [], []); + val empty = make_files (Path.current, [], [], []); fun extend _ = empty; fun merge _ = empty; ); fun map_files f = - Files.map (fn {master_dir, required, provided} => - make_files (f (master_dir, required, provided))); + Files.map (fn {master_dir, imports, required, provided} => + make_files (f (master_dir, imports, required, provided))); val master_directory = #master_dir o Files.get; +val imports_of = #imports o Files.get; -fun master dir = map_files (fn _ => (dir, [], [])); +fun put_deps dir imports = map_files (fn _ => (dir, imports, [], [])); fun require src_path = - map_files (fn (master_dir, required, provided) => + map_files (fn (master_dir, imports, required, provided) => if member (op =) required src_path then error ("Duplicate source file dependency: " ^ Path.implode src_path) - else (master_dir, src_path :: required, provided)); + else (master_dir, imports, src_path :: required, provided)); fun provide (src_path, path_id) = - map_files (fn (master_dir, required, provided) => + map_files (fn (master_dir, imports, required, provided) => if AList.defined (op =) provided src_path then error ("Duplicate resolution of source file dependency: " ^ Path.implode src_path) - else (master_dir, required, (src_path, path_id) :: provided)); + else (master_dir, imports, required, (src_path, path_id) :: provided)); (* maintain default paths *) @@ -251,9 +254,9 @@ (* begin theory *) -fun begin_theory dir name parents uses = +fun begin_theory dir name imports parents uses = Theory.begin_theory name parents - |> master dir + |> put_deps dir imports |> fold (require o fst) uses |> fold (fn (path, true) => Context.theory_map (exec_ml path) o Theory.checkpoint | _ => I) uses |> Theory.checkpoint; diff -r baf1964bc468 -r 621fa31e1dbd src/Pure/assumption.ML --- a/src/Pure/assumption.ML Mon Jan 24 22:29:50 2011 +0100 +++ b/src/Pure/assumption.ML Tue Jan 25 09:45:45 2011 +0100 @@ -79,10 +79,12 @@ fun extra_hyps ctxt th = subtract (op aconv) (map Thm.term_of (all_assms_of ctxt)) (Thm.hyps_of th); -(*named prems -- legacy feature*) val _ = Context.>> (Context.map_theory (Global_Theory.add_thms_dynamic (Binding.name "prems", - fn Context.Theory _ => [] | Context.Proof ctxt => all_prems_of ctxt))); + fn Context.Theory _ => [] + | Context.Proof ctxt => + (legacy_feature ("Use of global prems" ^ Position.str_of (Position.thread_data ())); + all_prems_of ctxt)))); (* local assumptions *) diff -r baf1964bc468 -r 621fa31e1dbd src/Tools/interpretation_with_defs.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/interpretation_with_defs.ML Tue Jan 25 09:45:45 2011 +0100 @@ -0,0 +1,96 @@ +(* Title: Tools/interpretation_with_defs.ML + Author: Florian Haftmann, TU Muenchen + +Interpretation accompanied with mixin definitions. EXPERIMENTAL. +*) + +signature INTERPRETATION_WITH_DEFS = +sig + val interpretation: Expression.expression_i -> + (Attrib.binding * ((binding * mixfix) * term)) list -> (Attrib.binding * term) list -> + theory -> Proof.state + val interpretation_cmd: Expression.expression -> + (Attrib.binding * ((binding * mixfix) * string)) list -> (Attrib.binding * string) list -> + theory -> Proof.state +end; + +structure Interpretation_With_Defs : INTERPRETATION_WITH_DEFS = +struct + +fun note_eqns_register deps witss def_eqns attrss eqns export export' context = + let + fun meta_rewrite context = + map (Local_Defs.meta_rewrite_rule (Context.proof_of context) #> Drule.abs_def) o + maps snd; + in + context + |> Element.generic_note_thmss Thm.lemmaK + (attrss ~~ map (fn eqn => [([Morphism.thm (export' $> export) eqn], [])]) eqns) + |-> (fn facts => `(fn context => meta_rewrite context facts)) + |-> (fn eqns => fold (fn ((dep, morph), wits) => + fn context => + Locale.add_registration (dep, morph $> Element.satisfy_morphism + (map (Element.morph_witness export') wits)) + (Element.eq_morphism (Context.theory_of context) (def_eqns @ eqns) |> + Option.map (rpair true)) + export context) (deps ~~ witss)) + end; + +local + +fun gen_interpretation prep_expr prep_decl parse_term parse_prop prep_attr + expression raw_defs raw_eqns theory = + let + val (_, (_, defs_ctxt)) = + prep_decl expression I [] (ProofContext.init_global theory); + + val rhss = map (parse_term defs_ctxt o snd o snd) raw_defs + |> Syntax.check_terms defs_ctxt; + val defs = map2 (fn (binding_thm, (binding_syn, _)) => fn rhs => + (binding_syn, (binding_thm, rhs))) raw_defs rhss; + + val (def_eqns, theory') = theory + |> Named_Target.theory_init + |> fold_map (Local_Theory.define) defs + |>> map (Thm.symmetric o snd o snd) + |> Local_Theory.exit_result_global (map o Morphism.thm); + + val ((propss, deps, export), expr_ctxt) = theory' + |> ProofContext.init_global + |> prep_expr expression; + + val eqns = map (parse_prop expr_ctxt o snd) raw_eqns + |> Syntax.check_terms expr_ctxt; + val attrss = map ((apsnd o map) (prep_attr theory) o fst) raw_eqns; + val goal_ctxt = fold Variable.auto_fixes eqns expr_ctxt; + val export' = Variable.export_morphism goal_ctxt expr_ctxt; + + fun after_qed witss eqns = + (ProofContext.background_theory o Context.theory_map) + (note_eqns_register deps witss def_eqns attrss eqns export export'); + + in Element.witness_proof_eqs after_qed propss eqns goal_ctxt end; + +in + +fun interpretation x = gen_interpretation Expression.cert_goal_expression + Expression.cert_declaration (K I) (K I) (K I) x; +fun interpretation_cmd x = gen_interpretation Expression.read_goal_expression + Expression.read_declaration Syntax.parse_term Syntax.parse_prop Attrib.intern_src x; + +end; + +val definesK = "defines"; +val _ = Keyword.keyword definesK; + +val _ = + Outer_Syntax.command "interpretation" + "prove interpretation of locale expression in theory" Keyword.thy_goal + (Parse.!!! (Parse_Spec.locale_expression true) -- + Scan.optional (Parse.$$$ definesK |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" + -- ((Parse.binding -- Parse.opt_mixfix') --| Parse.$$$ "is" -- Parse.term))) [] -- + Scan.optional (Parse.where_ |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)) [] + >> (fn ((expr, defs), equations) => Toplevel.print o + Toplevel.theory_to_proof (interpretation_cmd expr defs equations))); + +end; diff -r baf1964bc468 -r 621fa31e1dbd src/Tools/jEdit/plugin/Isabelle.props --- a/src/Tools/jEdit/plugin/Isabelle.props Mon Jan 24 22:29:50 2011 +0100 +++ b/src/Tools/jEdit/plugin/Isabelle.props Tue Jan 25 09:45:45 2011 +0100 @@ -33,7 +33,7 @@ options.isabelle.tooltip-margin=40 options.isabelle.tooltip-dismiss-delay.title=Tooltip Dismiss Delay (global) options.isabelle.tooltip-dismiss-delay=8.0 -options.isabelle.startup-timeout=10.0 +options.isabelle.startup-timeout=25.0 options.isabelle.auto-start.title=Auto Start options.isabelle.auto-start=true