# HG changeset patch # User chaieb # Date 1246713587 -7200 # Node ID f6d9798b13f1117ce0fe9d8519e397436d89fd88 # Parent 5fe21cac6bf7099e890d74e36cefc509c45261a9# Parent 36156921f5e5e7a6515f1d649be20cb44ac095c4 merged diff -r 36156921f5e5 -r f6d9798b13f1 .hgignore --- a/.hgignore Sat Jul 04 15:19:29 2009 +0200 +++ b/.hgignore Sat Jul 04 15:19:47 2009 +0200 @@ -17,9 +17,11 @@ ^doc-src/.*\.dvi ^doc-src/.*\.idx ^doc-src/.*\.ind +^doc-src/.*\.lof ^doc-src/.*\.log ^doc-src/.*\.out ^doc-src/.*\.rai ^doc-src/.*\.rao ^doc-src/.*\.toc + diff -r 36156921f5e5 -r f6d9798b13f1 Admin/MacOS/App1/README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Admin/MacOS/App1/README Sat Jul 04 15:19:47 2009 +0200 @@ -0,0 +1,9 @@ +Isabelle application bundle for MacOS +===================================== + +Requirements: + +* CocoaDialog 2.2.1 http://cocoadialog.sourceforge.net/ + +* Platypus 4.0 http://www.sveinbjorn.org/platypus + diff -r 36156921f5e5 -r f6d9798b13f1 Admin/MacOS/App1/mk --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Admin/MacOS/App1/mk Sat Jul 04 15:19:47 2009 +0200 @@ -0,0 +1,18 @@ +#!/bin/bash +# +# Make Isabelle application bundle + +THIS="$(cd "$(dirname "$0")"; pwd)" + +PLATYPUS_APP="/Applications/Platypus-4.0/Platypus.app" +COCOADIALOG_APP="/Applications/CocoaDialog.app" + +"$PLATYPUS_APP/Contents/Resources/platypus" \ + -a Isabelle -u Isabelle \ + -I "de.tum.in.isabelle" \ + -i "$THIS/../isabelle.icns" \ + -p /bin/bash \ + -c "$THIS/script" \ + -o None \ + -f "$COCOADIALOG_APP" \ + "$PWD/Isabelle.app" diff -r 36156921f5e5 -r f6d9798b13f1 Admin/MacOS/App1/script --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Admin/MacOS/App1/script Sat Jul 04 15:19:47 2009 +0200 @@ -0,0 +1,78 @@ +#!/bin/bash +# +# Author: Makarius +# +# Isabelle application wrapper + +THIS="$(cd "$(dirname "$0")"; pwd)" +THIS_APP="$(cd "$THIS/../.."; pwd)" +SUPER_APP="$(cd "$THIS/../../.."; pwd)" + + +# sane environment defaults +cd "$HOME" +PATH="$PATH:/opt/local/bin" + + +# settings support + +function choosefrom () +{ + local RESULT="" + local FILE="" + + for FILE in "$@" + do + [ -z "$RESULT" -a -e "$FILE" ] && RESULT="$FILE" + done + + [ -z "$RESULT" ] && RESULT="$FILE" + echo "$RESULT" +} + + +# Isabelle + +ISABELLE_TOOL="$(choosefrom \ + "$THIS/Isabelle/bin/isabelle" \ + "$SUPER_APP/Isabelle/bin/isabelle" \ + "$HOME/bin/isabelle" \ + isabelle)" + + +# Proof General / Emacs + +PROOFGENERAL_EMACS="$(choosefrom \ + "$THIS/Emacs.app/Contents/MacOS/Emacs" \ + "$SUPER_APP/Emacs.app/Contents/MacOS/Emacs" \ + /Applications/Emacs.app/Contents/MacOS/Emacs \ + "")" + +if [ -n "$PROOFGENERAL_EMACS" ]; then + EMACS_OPTIONS="-p $PROOFGENERAL_EMACS" +fi + + +# run interface with error feedback + +OUTPUT="/tmp/isabelle$$.out" + +( "$ISABELLE_TOOL" emacs $EMACS_OPTIONS "$@" ) > "$OUTPUT" 2>&1 +RC=$? + +if [ "$RC" != 0 ]; then + echo >> "$OUTPUT" + echo "Return code: $RC" >> "$OUTPUT" +fi + +if [ $(stat -f "%z" "$OUTPUT") != 0 ]; then + "$THIS/CocoaDialog.app/Contents/MacOS/CocoaDialog" textbox \ + --title "Isabelle" \ + --informative-text "Isabelle output" \ + --text-from-file "$OUTPUT" \ + --button1 "OK" +fi + +rm -f "$OUTPUT" + +exit "$RC" diff -r 36156921f5e5 -r f6d9798b13f1 Admin/MacOS/App2/Isabelle.app/Contents/Info.plist --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Admin/MacOS/App2/Isabelle.app/Contents/Info.plist Sat Jul 04 15:19:47 2009 +0200 @@ -0,0 +1,46 @@ + + + + + CFBundleDevelopmentRegion + English + CFBundleExecutable + Isabelle + CFBundleGetInfoString + Isabelle + CFBundleIconFile + isabelle.icns + CFBundleIdentifier + de.tum.in.isabelle + CFBundleInfoDictionaryVersion + 6.0 + CFBundleName + Isabelle + CFBundlePackageType + APPL + CFBundleShortVersionString + 2009 + CFBundleSignature + ???? + CFBundleVersion + 2009 + Java + + JVMVersion + 1.5+ + VMOptions + -Xmx384M + ClassPath + $JAVAROOT/isabelle-scala.jar + MainClass + isabelle.GUI_Setup + Properties + + isabelle.home + $APP_PACKAGE/Contents/Resources/Isabelle + apple.laf.useScreenMenuBar + true + + + + diff -r 36156921f5e5 -r f6d9798b13f1 Admin/MacOS/App2/Isabelle.app/Contents/MacOS/Isabelle --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Admin/MacOS/App2/Isabelle.app/Contents/MacOS/Isabelle Sat Jul 04 15:19:47 2009 +0200 @@ -0,0 +1,1 @@ +/System/Library/Frameworks/JavaVM.framework/Resources/MacOS/JavaApplicationStub \ No newline at end of file diff -r 36156921f5e5 -r f6d9798b13f1 Admin/MacOS/App2/README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Admin/MacOS/App2/README Sat Jul 04 15:19:47 2009 +0200 @@ -0,0 +1,7 @@ +Isabelle/JVM application bundle for MacOS +========================================= + +* http://developer.apple.com/documentation/Java/Conceptual/Java14Development/03-JavaDeployment/JavaDeployment.html + +* http://developer.apple.com/documentation/Java/Reference/Java_InfoplistRef/Articles/JavaDictionaryInfo.plistKeys.html#//apple_ref/doc/uid/TP40001969 + diff -r 36156921f5e5 -r f6d9798b13f1 Admin/MacOS/App2/mk --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Admin/MacOS/App2/mk Sat Jul 04 15:19:47 2009 +0200 @@ -0,0 +1,12 @@ +#!/bin/bash +# +# Make Isabelle/JVM application bundle + +THIS="$(cd "$(dirname "$0")"; pwd)" + +APP="$THIS/Isabelle.app" + +mkdir -p "$APP/Contents/Resources/Java" +cp "$THIS/../../../lib/classes/isabelle-scala.jar" "$APP/Contents/Resources/Java" +cp "$THIS/../isabelle.icns" "$APP/Contents/Resources" + diff -r 36156921f5e5 -r f6d9798b13f1 Admin/MacOS/README --- a/Admin/MacOS/README Sat Jul 04 15:19:29 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,9 +0,0 @@ -Isabelle application bundle for MacOS -===================================== - -Requirements: - -* CocoaDialog 2.2.1 http://cocoadialog.sourceforge.net/ - -* Platypus 4.0 http://www.sveinbjorn.org/platypus - diff -r 36156921f5e5 -r f6d9798b13f1 Admin/MacOS/mk --- a/Admin/MacOS/mk Sat Jul 04 15:19:29 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,18 +0,0 @@ -#!/bin/bash -# -# Make Isabelle application bundle - -THIS="$(cd "$(dirname "$0")"; pwd)" - -PLATYPUS_APP="/Applications/Platypus-4.0/Platypus.app" -COCOADIALOG_APP="/Applications/CocoaDialog.app" - -"$PLATYPUS_APP/Contents/Resources/platypus" \ - -a Isabelle -u Isabelle \ - -I "de.tum.in.isabelle" \ - -i "$THIS/isabelle.icns" \ - -p /bin/bash \ - -c "$THIS/script" \ - -o None \ - -f "$COCOADIALOG_APP" \ - "$PWD/Isabelle.app" diff -r 36156921f5e5 -r f6d9798b13f1 Admin/MacOS/script --- a/Admin/MacOS/script Sat Jul 04 15:19:29 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,78 +0,0 @@ -#!/bin/bash -# -# Author: Makarius -# -# Isabelle application wrapper - -THIS="$(cd "$(dirname "$0")"; pwd)" -THIS_APP="$(cd "$THIS/../.."; pwd)" -SUPER_APP="$(cd "$THIS/../../.."; pwd)" - - -# sane environment defaults -cd "$HOME" -PATH="$PATH:/opt/local/bin" - - -# settings support - -function choosefrom () -{ - local RESULT="" - local FILE="" - - for FILE in "$@" - do - [ -z "$RESULT" -a -e "$FILE" ] && RESULT="$FILE" - done - - [ -z "$RESULT" ] && RESULT="$FILE" - echo "$RESULT" -} - - -# Isabelle - -ISABELLE_TOOL="$(choosefrom \ - "$THIS/Isabelle/bin/isabelle" \ - "$SUPER_APP/Isabelle/bin/isabelle" \ - "$HOME/bin/isabelle" \ - isabelle)" - - -# Proof General / Emacs - -PROOFGENERAL_EMACS="$(choosefrom \ - "$THIS/Emacs.app/Contents/MacOS/Emacs" \ - "$SUPER_APP/Emacs.app/Contents/MacOS/Emacs" \ - /Applications/Emacs.app/Contents/MacOS/Emacs \ - "")" - -if [ -n "$PROOFGENERAL_EMACS" ]; then - EMACS_OPTIONS="-p $PROOFGENERAL_EMACS" -fi - - -# run interface with error feedback - -OUTPUT="/tmp/isabelle$$.out" - -( "$ISABELLE_TOOL" emacs $EMACS_OPTIONS "$@" ) > "$OUTPUT" 2>&1 -RC=$? - -if [ "$RC" != 0 ]; then - echo >> "$OUTPUT" - echo "Return code: $RC" >> "$OUTPUT" -fi - -if [ $(stat -f "%z" "$OUTPUT") != 0 ]; then - "$THIS/CocoaDialog.app/Contents/MacOS/CocoaDialog" textbox \ - --title "Isabelle" \ - --informative-text "Isabelle output" \ - --text-from-file "$OUTPUT" \ - --button1 "OK" -fi - -rm -f "$OUTPUT" - -exit "$RC" diff -r 36156921f5e5 -r f6d9798b13f1 Admin/launch4j/isabelle.xml --- a/Admin/launch4j/isabelle.xml Sat Jul 04 15:19:29 2009 +0200 +++ b/Admin/launch4j/isabelle.xml Sat Jul 04 15:19:47 2009 +0200 @@ -1,7 +1,7 @@ true gui - lib/classes/isabelle-scala.jar + lib\classes\isabelle-scala.jar Isabelle.exe @@ -10,7 +10,7 @@ http://java.com/download false - false + true isabelle.ico diff -r 36156921f5e5 -r f6d9798b13f1 Isabelle --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Isabelle Sat Jul 04 15:19:47 2009 +0200 @@ -0,0 +1,27 @@ +#!/usr/bin/env bash +# +# Author: Makarius +# +# Generic Isabelle application wrapper. + +if [ -L "$0" ]; then + TARGET="$(LC_ALL=C ls -l "$0" | sed 's/.* -> //')" + exec "$(cd "$(dirname "$0")"; cd "$(pwd -P)"; cd "$(dirname "$TARGET")"; pwd)/$(basename "$TARGET")" "$@" +fi + + +## settings + +ISABELLE_HOME="$(cd "$(dirname "$0")"; pwd -P)" +source "$ISABELLE_HOME/lib/scripts/getsettings" || exit 2 + +unset ISABELLE_SETTINGS_PRESENT + + +## main + +CLASSPATH="$(jvmpath "$CLASSPATH")" + +exec "$ISABELLE_JAVA" \ + "-Disabelle.home=$(jvmpath "$ISABELLE_HOME")" \ + -jar "$(jvmpath "$ISABELLE_HOME/lib/classes/isabelle-scala.jar")" "$@" diff -r 36156921f5e5 -r f6d9798b13f1 Isabelle.exe Binary file Isabelle.exe has changed diff -r 36156921f5e5 -r f6d9798b13f1 NEWS --- a/NEWS Sat Jul 04 15:19:29 2009 +0200 +++ b/NEWS Sat Jul 04 15:19:47 2009 +0200 @@ -37,11 +37,13 @@ * New method "linarith" invokes existing linear arithmetic decision procedure only. -* Implementation of quickcheck using generic code generator; default generators -are provided for all suitable HOL types, records and datatypes. - -* Constants Set.Pow and Set.image now with authentic syntax; object-logic definitions -Set.Pow_def and Set.image_def. INCOMPATIBILITY. +* Implementation of quickcheck using generic code generator; default +generators are provided for all suitable HOL types, records and +datatypes. + +* Constants Set.Pow and Set.image now with authentic syntax; +object-logic definitions Set.Pow_def and Set.image_def. +INCOMPATIBILITY. * Renamed theorems: Suc_eq_add_numeral_1 -> Suc_eq_plus1 @@ -49,10 +51,12 @@ Suc_plus1 -> Suc_eq_plus1 * New sledgehammer option "Full Types" in Proof General settings menu. -Causes full type information to be output to the ATPs. This slows ATPs down -considerably but eliminates a source of unsound "proofs" that fail later. - -* Discontinued ancient tradition to suffix certain ML module names with "_package", e.g.: +Causes full type information to be output to the ATPs. This slows +ATPs down considerably but eliminates a source of unsound "proofs" +that fail later. + +* Discontinued ancient tradition to suffix certain ML module names +with "_package", e.g.: DatatypePackage ~> Datatype InductivePackage ~> Inductive @@ -61,31 +65,35 @@ INCOMPATIBILITY. -* NewNumberTheory: Jeremy Avigad's new version of part of NumberTheory. -If possible, use NewNumberTheory, not NumberTheory. +* NewNumberTheory: Jeremy Avigad's new version of part of +NumberTheory. If possible, use NewNumberTheory, not NumberTheory. * Simplified interfaces of datatype module. INCOMPATIBILITY. -* Abbreviation "arbitrary" of "undefined" has disappeared; use "undefined" directly. -INCOMPATIBILITY. - -* New evaluator "approximate" approximates an real valued term using the same method as the -approximation method. - -* "approximate" supports now arithmetic expressions as boundaries of intervals and implements -interval splitting and taylor series expansion. - -* Changed DERIV_intros to a NamedThmsFun. Each of the theorems in DERIV_intros -assumes composition with an additional function and matches a variable to the -derivative, which has to be solved by the simplifier. Hence -(auto intro!: DERIV_intros) computes the derivative of most elementary terms. - -* Maclauren.DERIV_tac and Maclauren.deriv_tac was removed, they are replaced by: -(auto intro!: DERIV_intros) -INCOMPATIBILITY. +* Abbreviation "arbitrary" of "undefined" has disappeared; use +"undefined" directly. INCOMPATIBILITY. + +* New evaluator "approximate" approximates an real valued term using +the same method as the approximation method. + +* Method "approximate" supports now arithmetic expressions as +boundaries of intervals and implements interval splitting and Taylor +series expansion. + +* Changed DERIV_intros to a dynamic fact (via Named_Thms). Each of +the theorems in DERIV_intros assumes composition with an additional +function and matches a variable to the derivative, which has to be +solved by the simplifier. Hence (auto intro!: DERIV_intros) computes +the derivative of most elementary terms. + +* Maclauren.DERIV_tac and Maclauren.deriv_tac was removed, they are +replaced by: (auto intro!: DERIV_intros). INCOMPATIBILITY. + *** ML *** +* Renamed NamedThmsFun to Named_Thms. INCOMPATIBILITY. + * Eliminated old Attrib.add_attributes, Method.add_methods and related cominators for "args". INCOMPATIBILITY, need to use simplified Attrib/Method.setup introduced in Isabelle2009. diff -r 36156921f5e5 -r f6d9798b13f1 doc-src/Classes/Thy/Classes.thy --- a/doc-src/Classes/Thy/Classes.thy Sat Jul 04 15:19:29 2009 +0200 +++ b/doc-src/Classes/Thy/Classes.thy Sat Jul 04 15:19:47 2009 +0200 @@ -198,11 +198,11 @@ begin definition %quote - mult_prod_def: "p1 \ p2 = (fst p1 \ fst p2, snd p1 \ snd p2)" + mult_prod_def: "p\<^isub>1 \ p\<^isub>2 = (fst p\<^isub>1 \ fst p\<^isub>2, snd p\<^isub>1 \ snd p\<^isub>2)" instance %quote proof - fix p1 p2 p3 :: "\\semigroup \ \\semigroup" - show "p1 \ p2 \ p3 = p1 \ (p2 \ p3)" + fix p\<^isub>1 p\<^isub>2 p\<^isub>3 :: "\\semigroup \ \\semigroup" + show "p\<^isub>1 \ p\<^isub>2 \ p\<^isub>3 = p\<^isub>1 \ (p\<^isub>2 \ p\<^isub>3)" unfolding mult_prod_def by (simp add: assoc) qed diff -r 36156921f5e5 -r f6d9798b13f1 doc-src/Classes/Thy/document/Classes.tex --- a/doc-src/Classes/Thy/document/Classes.tex Sat Jul 04 15:19:29 2009 +0200 +++ b/doc-src/Classes/Thy/document/Classes.tex Sat Jul 04 15:19:47 2009 +0200 @@ -291,15 +291,15 @@ \isanewline \isacommand{definition}\isamarkupfalse% \isanewline -\ \ mult{\isacharunderscore}prod{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}p{\isadigit{1}}\ {\isasymotimes}\ p{\isadigit{2}}\ {\isacharequal}\ {\isacharparenleft}fst\ p{\isadigit{1}}\ {\isasymotimes}\ fst\ p{\isadigit{2}}{\isacharcomma}\ snd\ p{\isadigit{1}}\ {\isasymotimes}\ snd\ p{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ mult{\isacharunderscore}prod{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{2}}\ {\isacharequal}\ {\isacharparenleft}fst\ p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ fst\ p\isactrlisub {\isadigit{2}}{\isacharcomma}\ snd\ p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ snd\ p\isactrlisub {\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline \isanewline \isacommand{instance}\isamarkupfalse% \ \isacommand{proof}\isamarkupfalse% \isanewline \ \ \isacommand{fix}\isamarkupfalse% -\ p{\isadigit{1}}\ p{\isadigit{2}}\ p{\isadigit{3}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isasymColon}semigroup\ {\isasymtimes}\ {\isasymbeta}{\isasymColon}semigroup{\isachardoublequoteclose}\isanewline +\ p\isactrlisub {\isadigit{1}}\ p\isactrlisub {\isadigit{2}}\ p\isactrlisub {\isadigit{3}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isasymColon}semigroup\ {\isasymtimes}\ {\isasymbeta}{\isasymColon}semigroup{\isachardoublequoteclose}\isanewline \ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}p{\isadigit{1}}\ {\isasymotimes}\ p{\isadigit{2}}\ {\isasymotimes}\ p{\isadigit{3}}\ {\isacharequal}\ p{\isadigit{1}}\ {\isasymotimes}\ {\isacharparenleft}p{\isadigit{2}}\ {\isasymotimes}\ p{\isadigit{3}}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ {\isachardoublequoteopen}p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{2}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{3}}\ {\isacharequal}\ p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ {\isacharparenleft}p\isactrlisub {\isadigit{2}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{3}}{\isacharparenright}{\isachardoublequoteclose}\isanewline \ \ \ \ \isacommand{unfolding}\isamarkupfalse% \ mult{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% \ {\isacharparenleft}simp\ add{\isacharcolon}\ assoc{\isacharparenright}\isanewline diff -r 36156921f5e5 -r f6d9798b13f1 doc-src/IsarRef/Thy/HOL_Specific.thy --- a/doc-src/IsarRef/Thy/HOL_Specific.thy Sat Jul 04 15:19:29 2009 +0200 +++ b/doc-src/IsarRef/Thy/HOL_Specific.thy Sat Jul 04 15:19:47 2009 +0200 @@ -81,7 +81,7 @@ \end{matharray} \begin{rail} - 'split\_format' (((name *) + 'and') | ('(' 'complete' ')')) + 'split\_format' ((( name * ) + 'and') | ('(' 'complete' ')')) ; \end{rail} @@ -369,7 +369,7 @@ dtspec: parname? typespec infix? '=' (cons + '|') ; - cons: name (type *) mixfix? + cons: name ( type * ) mixfix? \end{rail} \begin{description} @@ -515,7 +515,7 @@ \begin{rail} 'relation' term ; - 'lexicographic\_order' (clasimpmod *) + 'lexicographic\_order' ( clasimpmod * ) ; \end{rail} @@ -565,7 +565,7 @@ ; recdeftc thmdecl? tc ; - hints: '(' 'hints' (recdefmod *) ')' + hints: '(' 'hints' ( recdefmod * ) ')' ; recdefmod: (('recdef\_simp' | 'recdef\_cong' | 'recdef\_wf') (() | 'add' | 'del') ':' thmrefs) | clasimpmod ; @@ -783,7 +783,7 @@ \end{matharray} \begin{rail} - 'iprover' ('!' ?) (rulemod *) + 'iprover' ('!' ?) ( rulemod * ) ; \end{rail} @@ -824,6 +824,74 @@ *} +section {* Checking and refuting propositions *} + +text {* + Identifying incorrect propositions usually involves evaluation of + particular assignments and systematic counter example search. This + is supported by the following commands. + + \begin{matharray}{rcl} + @{command_def (HOL) "value"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def (HOL) "quickcheck"}@{text "\<^sup>*"} & : & @{text "proof \"} \\ + @{command_def (HOL) "quickcheck_params"} & : & @{text "theory \ theory"} + \end{matharray} + + \begin{rail} + 'value' ( ( '[' name ']' ) ? ) modes? term + ; + + 'quickcheck' ( ( '[' args ']' ) ? ) nat? + ; + + 'quickcheck_params' ( ( '[' args ']' ) ? ) + ; + + modes: '(' (name + ) ')' + ; + + args: ( name '=' value + ',' ) + ; + \end{rail} + + \begin{description} + + \item @{command (HOL) "value"}~@{text t} evaluates and prints a + term; optionally @{text modes} can be specified, which are + appended to the current print mode (see also \cite{isabelle-ref}). + Internally, the evaluation is performed by registered evaluators, + which are invoked sequentially until a result is returned. + Alternatively a specific evaluator can be selected using square + brackets; available evaluators include @{text nbe} for + \emph{normalization by evaluation} and \emph{code} for code + generation in SML. + + \item @{command (HOL) "quickcheck"} tests the current goal for + counter examples using a series of arbitrary assignments for its + free variables; by default the first subgoal is tested, an other + can be selected explicitly using an optional goal index. + A number of configuration options are supported for + @{command (HOL) "quickcheck"}, notably: + + \begin{description} + + \item[size] specifies the maximum size of the search space for + assignment values. + + \item[iterations] sets how many sets of assignments are + generated for each particular size. + + \end{description} + + These option can be given within square brackets. + + \item @{command (HOL) "quickcheck_params"} changes quickcheck + configuration options persitently. + + \end{description} +*} + + section {* Invoking automated reasoning tools -- The Sledgehammer *} text {* @@ -860,7 +928,7 @@ \end{matharray} \begin{rail} - 'sledgehammer' (nameref *) + 'sledgehammer' ( nameref * ) ; 'atp\_messages' ('(' nat ')')? ; @@ -989,7 +1057,6 @@ (this actually covers the new-style theory format as well). \begin{matharray}{rcl} - @{command_def (HOL) "value"}@{text "\<^sup>*"} & : & @{text "context \"} \\ @{command_def (HOL) "code_module"} & : & @{text "theory \ theory"} \\ @{command_def (HOL) "code_library"} & : & @{text "theory \ theory"} \\ @{command_def (HOL) "consts_code"} & : & @{text "theory \ theory"} \\ @@ -998,9 +1065,6 @@ \end{matharray} \begin{rail} - 'value' term - ; - ( 'code\_module' | 'code\_library' ) modespec ? name ? \\ ( 'file' name ) ? ( 'imports' ( name + ) ) ? \\ 'contains' ( ( name '=' term ) + | term + ) @@ -1034,13 +1098,6 @@ ; \end{rail} - \begin{description} - - \item @{command (HOL) "value"}~@{text t} evaluates and prints a term - using the code generator. - - \end{description} - \medskip The other framework generates code from functional programs (including overloading using type classes) to SML \cite{SML}, OCaml \cite{OCaml} and Haskell \cite{haskell-revised-report}. diff -r 36156921f5e5 -r f6d9798b13f1 doc-src/IsarRef/Thy/Spec.thy --- a/doc-src/IsarRef/Thy/Spec.thy Sat Jul 04 15:19:29 2009 +0200 +++ b/doc-src/IsarRef/Thy/Spec.thy Sat Jul 04 15:19:47 2009 +0200 @@ -600,7 +600,7 @@ ; 'instance' ; - 'instance' nameref '::' arity + 'instance' (nameref + 'and') '::' arity ; 'subclass' target? nameref ; @@ -644,7 +644,7 @@ concluded by an @{command_ref (local) "end"} command. Note that a list of simultaneous type constructors may be given; - this corresponds nicely to mutual recursive type definitions, e.g.\ + this corresponds nicely to mutually recursive type definitions, e.g.\ in Isabelle/HOL. \item @{command "instance"} in an instantiation target body sets diff -r 36156921f5e5 -r f6d9798b13f1 doc-src/IsarRef/Thy/document/HOL_Specific.tex --- a/doc-src/IsarRef/Thy/document/HOL_Specific.tex Sat Jul 04 15:19:29 2009 +0200 +++ b/doc-src/IsarRef/Thy/document/HOL_Specific.tex Sat Jul 04 15:19:47 2009 +0200 @@ -98,7 +98,7 @@ \end{matharray} \begin{rail} - 'split\_format' (((name *) + 'and') | ('(' 'complete' ')')) + 'split\_format' ((( name * ) + 'and') | ('(' 'complete' ')')) ; \end{rail} @@ -374,7 +374,7 @@ dtspec: parname? typespec infix? '=' (cons + '|') ; - cons: name (type *) mixfix? + cons: name ( type * ) mixfix? \end{rail} \begin{description} @@ -520,7 +520,7 @@ \begin{rail} 'relation' term ; - 'lexicographic\_order' (clasimpmod *) + 'lexicographic\_order' ( clasimpmod * ) ; \end{rail} @@ -570,7 +570,7 @@ ; recdeftc thmdecl? tc ; - hints: '(' 'hints' (recdefmod *) ')' + hints: '(' 'hints' ( recdefmod * ) ')' ; recdefmod: (('recdef\_simp' | 'recdef\_cong' | 'recdef\_wf') (() | 'add' | 'del') ':' thmrefs) | clasimpmod ; @@ -793,7 +793,7 @@ \end{matharray} \begin{rail} - 'iprover' ('!' ?) (rulemod *) + 'iprover' ('!' ?) ( rulemod * ) ; \end{rail} @@ -835,6 +835,76 @@ \end{isamarkuptext}% \isamarkuptrue% % +\isamarkupsection{Checking and refuting propositions% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Identifying incorrect propositions usually involves evaluation of + particular assignments and systematic counter example search. This + is supported by the following commands. + + \begin{matharray}{rcl} + \indexdef{HOL}{command}{value}\hypertarget{command.HOL.value}{\hyperlink{command.HOL.value}{\mbox{\isa{\isacommand{value}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\ + \indexdef{HOL}{command}{quickcheck}\hypertarget{command.HOL.quickcheck}{\hyperlink{command.HOL.quickcheck}{\mbox{\isa{\isacommand{quickcheck}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}proof\ {\isasymrightarrow}{\isachardoublequote}} \\ + \indexdef{HOL}{command}{quickcheck\_params}\hypertarget{command.HOL.quickcheck-params}{\hyperlink{command.HOL.quickcheck-params}{\mbox{\isa{\isacommand{quickcheck{\isacharunderscore}params}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} + \end{matharray} + + \begin{rail} + 'value' ( ( '[' name ']' ) ? ) modes? term + ; + + 'quickcheck' ( ( '[' args ']' ) ? ) nat? + ; + + 'quickcheck_params' ( ( '[' args ']' ) ? ) + ; + + modes: '(' (name + ) ')' + ; + + args: ( name '=' value + ',' ) + ; + \end{rail} + + \begin{description} + + \item \hyperlink{command.HOL.value}{\mbox{\isa{\isacommand{value}}}}~\isa{t} evaluates and prints a + term; optionally \isa{modes} can be specified, which are + appended to the current print mode (see also \cite{isabelle-ref}). + Internally, the evaluation is performed by registered evaluators, + which are invoked sequentially until a result is returned. + Alternatively a specific evaluator can be selected using square + brackets; available evaluators include \isa{nbe} for + \emph{normalization by evaluation} and \emph{code} for code + generation in SML. + + \item \hyperlink{command.HOL.quickcheck}{\mbox{\isa{\isacommand{quickcheck}}}} tests the current goal for + counter examples using a series of arbitrary assignments for its + free variables; by default the first subgoal is tested, an other + can be selected explicitly using an optional goal index. + A number of configuration options are supported for + \hyperlink{command.HOL.quickcheck}{\mbox{\isa{\isacommand{quickcheck}}}}, notably: + + \begin{description} + + \item[size] specifies the maximum size of the search space for + assignment values. + + \item[iterations] sets how many sets of assignments are + generated for each particular size. + + \end{description} + + These option can be given within square brackets. + + \item \hyperlink{command.HOL.quickcheck-params}{\mbox{\isa{\isacommand{quickcheck{\isacharunderscore}params}}}} changes quickcheck + configuration options persitently. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% \isamarkupsection{Invoking automated reasoning tools -- The Sledgehammer% } \isamarkuptrue% @@ -873,7 +943,7 @@ \end{matharray} \begin{rail} - 'sledgehammer' (nameref *) + 'sledgehammer' ( nameref * ) ; 'atp\_messages' ('(' nat ')')? ; @@ -998,7 +1068,6 @@ (this actually covers the new-style theory format as well). \begin{matharray}{rcl} - \indexdef{HOL}{command}{value}\hypertarget{command.HOL.value}{\hyperlink{command.HOL.value}{\mbox{\isa{\isacommand{value}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\ \indexdef{HOL}{command}{code\_module}\hypertarget{command.HOL.code-module}{\hyperlink{command.HOL.code-module}{\mbox{\isa{\isacommand{code{\isacharunderscore}module}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\ \indexdef{HOL}{command}{code\_library}\hypertarget{command.HOL.code-library}{\hyperlink{command.HOL.code-library}{\mbox{\isa{\isacommand{code{\isacharunderscore}library}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\ \indexdef{HOL}{command}{consts\_code}\hypertarget{command.HOL.consts-code}{\hyperlink{command.HOL.consts-code}{\mbox{\isa{\isacommand{consts{\isacharunderscore}code}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\ @@ -1007,9 +1076,6 @@ \end{matharray} \begin{rail} - 'value' term - ; - ( 'code\_module' | 'code\_library' ) modespec ? name ? \\ ( 'file' name ) ? ( 'imports' ( name + ) ) ? \\ 'contains' ( ( name '=' term ) + | term + ) @@ -1043,13 +1109,6 @@ ; \end{rail} - \begin{description} - - \item \hyperlink{command.HOL.value}{\mbox{\isa{\isacommand{value}}}}~\isa{t} evaluates and prints a term - using the code generator. - - \end{description} - \medskip The other framework generates code from functional programs (including overloading using type classes) to SML \cite{SML}, OCaml \cite{OCaml} and Haskell \cite{haskell-revised-report}. diff -r 36156921f5e5 -r f6d9798b13f1 doc-src/IsarRef/Thy/document/Spec.tex --- a/doc-src/IsarRef/Thy/document/Spec.tex Sat Jul 04 15:19:29 2009 +0200 +++ b/doc-src/IsarRef/Thy/document/Spec.tex Sat Jul 04 15:19:47 2009 +0200 @@ -614,7 +614,7 @@ ; 'instance' ; - 'instance' nameref '::' arity + 'instance' (nameref + 'and') '::' arity ; 'subclass' target? nameref ; @@ -653,7 +653,7 @@ concluded by an \indexref{local}{command}{end}\hyperlink{command.local.end}{\mbox{\isa{\isacommand{end}}}} command. Note that a list of simultaneous type constructors may be given; - this corresponds nicely to mutual recursive type definitions, e.g.\ + this corresponds nicely to mutually recursive type definitions, e.g.\ in Isabelle/HOL. \item \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} in an instantiation target body sets @@ -733,7 +733,7 @@ \end{matharray} Axiomatic type classes are Isabelle/Pure's primitive - \emph{definitional} interface to type classes. For practical + interface to type classes. For practical applications, you should consider using classes (cf.~\secref{sec:classes}) which provide high level interface. diff -r 36156921f5e5 -r f6d9798b13f1 etc/settings --- a/etc/settings Sat Jul 04 15:19:29 2009 +0200 +++ b/etc/settings Sat Jul 04 15:19:47 2009 +0200 @@ -66,14 +66,15 @@ ### JVM components (Scala or Java) ### +ISABELLE_JAVA="java" ISABELLE_SCALA="scala" -ISABELLE_JAVA="java" -if [ -e "$ISABELLE_HOME/contrib/scala" ]; then - classpath "$ISABELLE_HOME/contrib/scala/lib/scala-library.jar" -elif [ -e "$ISABELLE_HOME/../scala" ]; then - classpath "$ISABELLE_HOME/../scala/lib/scala-library.jar" -fi +[ -z "$SCALA_HOME" ] && SCALA_HOME=$(choosefrom \ + "$ISABELLE_HOME/contrib/scala" \ + "$ISABELLE_HOME/../scala" \ + "") + +[ -n "$SCALA_HOME" ] && ISABELLE_SCALA="$SCALA_HOME/bin/scala" classpath "$ISABELLE_HOME/lib/classes/Pure.jar" diff -r 36156921f5e5 -r f6d9798b13f1 src/CCL/Wfd.thy --- a/src/CCL/Wfd.thy Sat Jul 04 15:19:29 2009 +0200 +++ b/src/CCL/Wfd.thy Sat Jul 04 15:19:47 2009 +0200 @@ -495,7 +495,7 @@ ML {* local - structure Data = NamedThmsFun(val name = "eval" val description = "evaluation rules"); + structure Data = Named_Thms(val name = "eval" val description = "evaluation rules"); in fun eval_tac ctxt ths = diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Deriv.thy Sat Jul 04 15:19:47 2009 +0200 @@ -264,21 +264,23 @@ by (simp add: d dfx real_scaleR_def) qed -(* let's do the standard proof though theorem *) -(* LIM_mult2 follows from a NS proof *) +text {* + Let's do the standard proof, though theorem + @{text "LIM_mult2"} follows from a NS proof +*} lemma DERIV_cmult: "DERIV f x :> D ==> DERIV (%x. c * f x) x :> c*D" by (drule DERIV_mult' [OF DERIV_const], simp) -(* standard version *) +text {* Standard version *} lemma DERIV_chain: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (f o g) x :> Da * Db" by (drule (1) DERIV_chain', simp add: o_def real_scaleR_def mult_commute) lemma DERIV_chain2: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (%x. f (g x)) x :> Da * Db" by (auto dest: DERIV_chain simp add: o_def) -(*derivative of linear multiplication*) +text {* Derivative of linear multiplication *} lemma DERIV_cmult_Id [simp]: "DERIV (op * c) x :> c" by (cut_tac c = c and x = x in DERIV_ident [THEN DERIV_cmult], simp) @@ -287,21 +289,21 @@ apply (simp add: real_scaleR_def real_of_nat_def) done -text{*Power of -1*} +text {* Power of @{text "-1"} *} lemma DERIV_inverse: fixes x :: "'a::{real_normed_field}" shows "x \ 0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ Suc (Suc 0)))" by (drule DERIV_inverse' [OF DERIV_ident]) simp -text{*Derivative of inverse*} +text {* Derivative of inverse *} lemma DERIV_inverse_fun: fixes x :: "'a::{real_normed_field}" shows "[| DERIV f x :> d; f(x) \ 0 |] ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))" by (drule (1) DERIV_inverse') (simp add: mult_ac nonzero_inverse_mult_distrib) -text{*Derivative of quotient*} +text {* Derivative of quotient *} lemma DERIV_quotient: fixes x :: "'a::{real_normed_field}" shows "[| DERIV f x :> d; DERIV g x :> e; g(x) \ 0 |] @@ -311,14 +313,16 @@ lemma lemma_DERIV_subst: "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E" by auto -text {* DERIV_intros *} +text {* @{text "DERIV_intros"} *} +ML {* +structure Deriv_Intros = Named_Thms +( + val name = "DERIV_intros" + val description = "DERIV introduction rules" +) +*} -ML{* -structure DerivIntros = - NamedThmsFun(val name = "DERIV_intros" - val description = "DERIV introduction rules"); -*} -setup DerivIntros.setup +setup Deriv_Intros.setup lemma DERIV_cong: "\ DERIV f x :> X ; X = Y \ \ DERIV f x :> Y" by simp @@ -336,6 +340,7 @@ unfolded real_of_nat_def[symmetric], DERIV_intros] DERIV_setsum[THEN DERIV_cong, DERIV_intros] + subsection {* Differentiability predicate *} definition diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Finite_Set.thy Sat Jul 04 15:19:47 2009 +0200 @@ -1883,14 +1883,14 @@ (if a:A then setprod f A / f a else setprod f A)" by (erule finite_induct) (auto simp add: insert_Diff_if) -lemma setprod_inversef: "finite A ==> - ALL x: A. f x \ (0::'a::{field,division_by_zero}) ==> - setprod (inverse \ f) A = inverse (setprod f A)" +lemma setprod_inversef: + fixes f :: "'b \ 'a::{field,division_by_zero}" + shows "finite A ==> setprod (inverse \ f) A = inverse (setprod f A)" by (erule finite_induct) auto lemma setprod_dividef: - "[|finite A; - \x \ A. g x \ (0::'a::{field,division_by_zero})|] + fixes f :: "'b \ 'a::{field,division_by_zero}" + shows "finite A ==> setprod (%x. f x / g x) A = setprod f A / setprod g A" apply (subgoal_tac "setprod (%x. f x / g x) A = setprod (%x. f x * (inverse \ g) x) A") diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/HOL.thy --- a/src/HOL/HOL.thy Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/HOL.thy Sat Jul 04 15:19:47 2009 +0200 @@ -923,9 +923,11 @@ ML_Antiquote.value "claset" (Scan.succeed "Classical.local_claset_of (ML_Context.the_local_context ())"); -structure ResAtpset = NamedThmsFun(val name = "atp" val description = "ATP rules"); +structure ResAtpset = Named_Thms + (val name = "atp" val description = "ATP rules"); -structure ResBlacklist = NamedThmsFun(val name = "noatp" val description = "theorems blacklisted for ATP"); +structure ResBlacklist = Named_Thms + (val name = "noatp" val description = "theorems blacklisted for ATP"); *} text {*ResBlacklist holds theorems blacklisted to sledgehammer. @@ -1982,19 +1984,18 @@ Scan.succeed (K (SIMPLE_METHOD' (CONVERSION Nbe.norm_conv THEN' (fn k => TRY (rtac TrueI k))))) *} "solve goal by normalization" + subsubsection {* Quickcheck *} ML {* -structure Quickcheck_RecFun_Simp_Thms = NamedThmsFun +structure Quickcheck_RecFun_Simps = Named_Thms ( val name = "quickcheck_recfun_simp" val description = "simplification rules of recursive functions as needed by Quickcheck" ) *} -setup {* - Quickcheck_RecFun_Simp_Thms.setup -*} +setup Quickcheck_RecFun_Simps.setup setup {* Quickcheck.add_generator ("SML", Codegen.test_term) @@ -2008,22 +2009,22 @@ text {* This will be relocated once Nitpick is moved to HOL. *} ML {* -structure Nitpick_Const_Def_Thms = NamedThmsFun +structure Nitpick_Const_Defs = Named_Thms ( val name = "nitpick_const_def" val description = "alternative definitions of constants as needed by Nitpick" ) -structure Nitpick_Const_Simp_Thms = NamedThmsFun +structure Nitpick_Const_Simps = Named_Thms ( val name = "nitpick_const_simp" val description = "equational specification of constants as needed by Nitpick" ) -structure Nitpick_Const_Psimp_Thms = NamedThmsFun +structure Nitpick_Const_Psimps = Named_Thms ( val name = "nitpick_const_psimp" val description = "partial equational specification of constants as needed by Nitpick" ) -structure Nitpick_Ind_Intro_Thms = NamedThmsFun +structure Nitpick_Ind_Intros = Named_Thms ( val name = "nitpick_ind_intro" val description = "introduction rules for (co)inductive predicates as needed by Nitpick" @@ -2031,10 +2032,10 @@ *} setup {* - Nitpick_Const_Def_Thms.setup - #> Nitpick_Const_Simp_Thms.setup - #> Nitpick_Const_Psimp_Thms.setup - #> Nitpick_Ind_Intro_Thms.setup + Nitpick_Const_Defs.setup + #> Nitpick_Const_Simps.setup + #> Nitpick_Const_Psimps.setup + #> Nitpick_Ind_Intros.setup *} diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/IsaMakefile Sat Jul 04 15:19:47 2009 +0200 @@ -1003,11 +1003,11 @@ $(OUT)/HOL-Nominal: $(OUT)/HOL Nominal/ROOT.ML \ Nominal/Nominal.thy \ Nominal/nominal_atoms.ML \ + Nominal/nominal_datatype.ML \ Nominal/nominal_fresh_fun.ML \ Nominal/nominal_induct.ML \ Nominal/nominal_inductive.ML \ Nominal/nominal_inductive2.ML \ - Nominal/nominal.ML \ Nominal/nominal_permeq.ML \ Nominal/nominal_primrec.ML \ Nominal/nominal_thmdecls.ML \ diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Library/Fset.thy --- a/src/HOL/Library/Fset.thy Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Library/Fset.thy Sat Jul 04 15:19:47 2009 +0200 @@ -7,11 +7,6 @@ imports List_Set begin -lemma foldl_apply_inv: - assumes "\x. g (h x) = x" - shows "foldl f (g s) xs = g (foldl (\s x. h (f (g s) x)) s xs)" - by (rule sym, induct xs arbitrary: s) (simp_all add: assms) - declare mem_def [simp] diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Limits.thy --- a/src/HOL/Limits.thy Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Limits.thy Sat Jul 04 15:19:47 2009 +0200 @@ -350,7 +350,7 @@ lemmas Zfun_mult_left = mult.Zfun_left -subsection{* Limits *} +subsection {* Limits *} definition tendsto :: "('a \ 'b::topological_space) \ 'b \ 'a net \ bool" @@ -358,13 +358,15 @@ where [code del]: "(f ---> l) net \ (\S. open S \ l \ S \ eventually (\x. f x \ S) net)" -ML{* -structure TendstoIntros = - NamedThmsFun(val name = "tendsto_intros" - val description = "introduction rules for tendsto"); +ML {* +structure Tendsto_Intros = Named_Thms +( + val name = "tendsto_intros" + val description = "introduction rules for tendsto" +) *} -setup TendstoIntros.setup +setup Tendsto_Intros.setup lemma topological_tendstoI: "(\S. open S \ l \ S \ eventually (\x. f x \ S) net) diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/List.thy --- a/src/HOL/List.thy Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/List.thy Sat Jul 04 15:19:47 2009 +0200 @@ -2080,6 +2080,11 @@ "foldl g a (map f xs) = foldl (%a x. g a (f x)) a xs" by(induct xs arbitrary:a) simp_all +lemma foldl_apply_inv: + assumes "\x. g (h x) = x" + shows "foldl f (g s) xs = g (foldl (\s x. h (f (g s) x)) s xs)" + by (rule sym, induct xs arbitrary: s) (simp_all add: assms) + lemma foldl_cong [fundef_cong, recdef_cong]: "[| a = b; l = k; !!a x. x : set l ==> f a x = g a x |] ==> foldl f a l = foldl g b k" diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Nominal/Nominal.thy --- a/src/HOL/Nominal/Nominal.thy Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Nominal/Nominal.thy Sat Jul 04 15:19:47 2009 +0200 @@ -3,7 +3,7 @@ uses ("nominal_thmdecls.ML") ("nominal_atoms.ML") - ("nominal.ML") + ("nominal_datatype.ML") ("nominal_induct.ML") ("nominal_permeq.ML") ("nominal_fresh_fun.ML") @@ -3670,7 +3670,7 @@ lemma allE_Nil: assumes "\x. P x" obtains "P []" using assms .. -use "nominal.ML" +use "nominal_datatype.ML" (******************************************************) (* primitive recursive functions on nominal datatypes *) diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Nominal/nominal.ML --- a/src/HOL/Nominal/nominal.ML Sat Jul 04 15:19:29 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2094 +0,0 @@ -(* Title: HOL/Nominal/nominal.ML - Author: Stefan Berghofer and Christian Urban, TU Muenchen - -Nominal datatype package for Isabelle/HOL. -*) - -signature NOMINAL = -sig - val add_nominal_datatype : Datatype.config -> string list -> - (string list * bstring * mixfix * - (bstring * string list * mixfix) list) list -> theory -> theory - type descr - type nominal_datatype_info - val get_nominal_datatypes : theory -> nominal_datatype_info Symtab.table - val get_nominal_datatype : theory -> string -> nominal_datatype_info option - val mk_perm: typ list -> term -> term -> term - val perm_of_pair: term * term -> term - val mk_not_sym: thm list -> thm list - val perm_simproc: simproc - val fresh_const: typ -> typ -> term - val fresh_star_const: typ -> typ -> term -end - -structure Nominal : NOMINAL = -struct - -val finite_emptyI = thm "finite.emptyI"; -val finite_Diff = thm "finite_Diff"; -val finite_Un = thm "finite_Un"; -val Un_iff = thm "Un_iff"; -val In0_eq = thm "In0_eq"; -val In1_eq = thm "In1_eq"; -val In0_not_In1 = thm "In0_not_In1"; -val In1_not_In0 = thm "In1_not_In0"; -val Un_assoc = thm "Un_assoc"; -val Collect_disj_eq = thm "Collect_disj_eq"; -val empty_def = thm "empty_def"; -val empty_iff = thm "empty_iff"; - -open DatatypeAux; -open NominalAtoms; - -(** FIXME: Datatype should export this function **) - -local - -fun dt_recs (DtTFree _) = [] - | dt_recs (DtType (_, dts)) = List.concat (map dt_recs dts) - | dt_recs (DtRec i) = [i]; - -fun dt_cases (descr: descr) (_, args, constrs) = - let - fun the_bname i = Long_Name.base_name (#1 (valOf (AList.lookup (op =) descr i))); - val bnames = map the_bname (distinct op = (List.concat (map dt_recs args))); - in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end; - - -fun induct_cases descr = - DatatypeProp.indexify_names (List.concat (map (dt_cases descr) (map #2 descr))); - -fun exhaust_cases descr i = dt_cases descr (valOf (AList.lookup (op =) descr i)); - -in - -fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr); - -fun mk_case_names_exhausts descr new = - map (RuleCases.case_names o exhaust_cases descr o #1) - (List.filter (fn ((_, (name, _, _))) => name mem_string new) descr); - -end; - -(* theory data *) - -type descr = (int * (string * dtyp list * (string * (dtyp list * dtyp) list) list)) list; - -type nominal_datatype_info = - {index : int, - descr : descr, - sorts : (string * sort) list, - rec_names : string list, - rec_rewrites : thm list, - induction : thm, - distinct : thm list, - inject : thm list}; - -structure NominalDatatypesData = TheoryDataFun -( - type T = nominal_datatype_info Symtab.table; - val empty = Symtab.empty; - val copy = I; - val extend = I; - fun merge _ tabs : T = Symtab.merge (K true) tabs; -); - -val get_nominal_datatypes = NominalDatatypesData.get; -val put_nominal_datatypes = NominalDatatypesData.put; -val map_nominal_datatypes = NominalDatatypesData.map; -val get_nominal_datatype = Symtab.lookup o get_nominal_datatypes; - - -(**** make datatype info ****) - -fun make_dt_info descr sorts induct reccomb_names rec_thms - (((i, (_, (tname, _, _))), distinct), inject) = - (tname, - {index = i, - descr = descr, - sorts = sorts, - rec_names = reccomb_names, - rec_rewrites = rec_thms, - induction = induct, - distinct = distinct, - inject = inject}); - -(*******************************) - -val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma); - - -(** simplification procedure for sorting permutations **) - -val dj_cp = thm "dj_cp"; - -fun dest_permT (Type ("fun", [Type ("List.list", [Type ("*", [T, _])]), - Type ("fun", [_, U])])) = (T, U); - -fun permTs_of (Const ("Nominal.perm", T) $ t $ u) = fst (dest_permT T) :: permTs_of u - | permTs_of _ = []; - -fun perm_simproc' thy ss (Const ("Nominal.perm", T) $ t $ (u as Const ("Nominal.perm", U) $ r $ s)) = - let - val (aT as Type (a, []), S) = dest_permT T; - val (bT as Type (b, []), _) = dest_permT U - in if aT mem permTs_of u andalso aT <> bT then - let - val cp = cp_inst_of thy a b; - val dj = dj_thm_of thy b a; - val dj_cp' = [cp, dj] MRS dj_cp; - val cert = SOME o cterm_of thy - in - SOME (mk_meta_eq (Drule.instantiate' [SOME (ctyp_of thy S)] - [cert t, cert r, cert s] dj_cp')) - end - else NONE - end - | perm_simproc' thy ss _ = NONE; - -val perm_simproc = - Simplifier.simproc (the_context ()) "perm_simp" ["pi1 \ (pi2 \ x)"] perm_simproc'; - -val meta_spec = thm "meta_spec"; - -fun projections rule = - ProjectRule.projections (ProofContext.init (Thm.theory_of_thm rule)) rule - |> map (standard #> RuleCases.save rule); - -val supp_prod = thm "supp_prod"; -val fresh_prod = thm "fresh_prod"; -val supports_fresh = thm "supports_fresh"; -val supports_def = thm "Nominal.supports_def"; -val fresh_def = thm "fresh_def"; -val supp_def = thm "supp_def"; -val rev_simps = thms "rev.simps"; -val app_simps = thms "append.simps"; -val at_fin_set_supp = thm "at_fin_set_supp"; -val at_fin_set_fresh = thm "at_fin_set_fresh"; -val abs_fun_eq1 = thm "abs_fun_eq1"; - -val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq]; - -fun mk_perm Ts t u = - let - val T = fastype_of1 (Ts, t); - val U = fastype_of1 (Ts, u) - in Const ("Nominal.perm", T --> U --> U) $ t $ u end; - -fun perm_of_pair (x, y) = - let - val T = fastype_of x; - val pT = mk_permT T - in Const ("List.list.Cons", HOLogic.mk_prodT (T, T) --> pT --> pT) $ - HOLogic.mk_prod (x, y) $ Const ("List.list.Nil", pT) - end; - -fun mk_not_sym ths = maps (fn th => case prop_of th of - _ $ (Const ("Not", _) $ (Const ("op =", _) $ _ $ _)) => [th, th RS not_sym] - | _ => [th]) ths; - -fun fresh_const T U = Const ("Nominal.fresh", T --> U --> HOLogic.boolT); -fun fresh_star_const T U = - Const ("Nominal.fresh_star", HOLogic.mk_setT T --> U --> HOLogic.boolT); - -fun gen_add_nominal_datatype prep_typ config new_type_names dts thy = - let - (* this theory is used just for parsing *) - - val tmp_thy = thy |> - Theory.copy |> - Sign.add_types (map (fn (tvs, tname, mx, _) => - (Binding.name tname, length tvs, mx)) dts); - - val atoms = atoms_of thy; - - fun prep_constr ((constrs, sorts), (cname, cargs, mx)) = - let val (cargs', sorts') = Library.foldl (prep_typ tmp_thy) (([], sorts), cargs) - in (constrs @ [(cname, cargs', mx)], sorts') end - - fun prep_dt_spec ((dts, sorts), (tvs, tname, mx, constrs)) = - let val (constrs', sorts') = Library.foldl prep_constr (([], sorts), constrs) - in (dts @ [(tvs, tname, mx, constrs')], sorts') end - - val (dts', sorts) = Library.foldl prep_dt_spec (([], []), dts); - val tyvars = map (map (fn s => - (s, the (AList.lookup (op =) sorts s))) o #1) dts'; - - fun inter_sort thy S S' = Type.inter_sort (Sign.tsig_of thy) (S, S'); - fun augment_sort_typ thy S = - let val S = Sign.certify_sort thy S - in map_type_tfree (fn (s, S') => TFree (s, - if member (op = o apsnd fst) sorts s then inter_sort thy S S' else S')) - end; - fun augment_sort thy S = map_types (augment_sort_typ thy S); - - val types_syntax = map (fn (tvs, tname, mx, constrs) => (tname, mx)) dts'; - val constr_syntax = map (fn (tvs, tname, mx, constrs) => - map (fn (cname, cargs, mx) => (cname, mx)) constrs) dts'; - - val ps = map (fn (_, n, _, _) => - (Sign.full_bname tmp_thy n, Sign.full_bname tmp_thy (n ^ "_Rep"))) dts; - val rps = map Library.swap ps; - - fun replace_types (Type ("Nominal.ABS", [T, U])) = - Type ("fun", [T, Type ("Nominal.noption", [replace_types U])]) - | replace_types (Type (s, Ts)) = - Type (getOpt (AList.lookup op = ps s, s), map replace_types Ts) - | replace_types T = T; - - val dts'' = map (fn (tvs, tname, mx, constrs) => (tvs, Binding.name (tname ^ "_Rep"), NoSyn, - map (fn (cname, cargs, mx) => (Binding.name (cname ^ "_Rep"), - map replace_types cargs, NoSyn)) constrs)) dts'; - - val new_type_names' = map (fn n => n ^ "_Rep") new_type_names; - - val (full_new_type_names',thy1) = - Datatype.add_datatype config new_type_names' dts'' thy; - - val {descr, induction, ...} = - Datatype.the_info thy1 (hd full_new_type_names'); - fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); - - val big_name = space_implode "_" new_type_names; - - - (**** define permutation functions ****) - - val permT = mk_permT (TFree ("'x", HOLogic.typeS)); - val pi = Free ("pi", permT); - val perm_types = map (fn (i, _) => - let val T = nth_dtyp i - in permT --> T --> T end) descr; - val perm_names' = DatatypeProp.indexify_names (map (fn (i, _) => - "perm_" ^ name_of_typ (nth_dtyp i)) descr); - val perm_names = replicate (length new_type_names) "Nominal.perm" @ - map (Sign.full_bname thy1) (List.drop (perm_names', length new_type_names)); - val perm_names_types = perm_names ~~ perm_types; - val perm_names_types' = perm_names' ~~ perm_types; - - val perm_eqs = maps (fn (i, (_, _, constrs)) => - let val T = nth_dtyp i - in map (fn (cname, dts) => - let - val Ts = map (typ_of_dtyp descr sorts) dts; - val names = Name.variant_list ["pi"] (DatatypeProp.make_tnames Ts); - val args = map Free (names ~~ Ts); - val c = Const (cname, Ts ---> T); - fun perm_arg (dt, x) = - let val T = type_of x - in if is_rec_type dt then - let val (Us, _) = strip_type T - in list_abs (map (pair "x") Us, - Free (nth perm_names_types' (body_index dt)) $ pi $ - list_comb (x, map (fn (i, U) => - Const ("Nominal.perm", permT --> U --> U) $ - (Const ("List.rev", permT --> permT) $ pi) $ - Bound i) ((length Us - 1 downto 0) ~~ Us))) - end - else Const ("Nominal.perm", permT --> T --> T) $ pi $ x - end; - in - (Attrib.empty_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq - (Free (nth perm_names_types' i) $ - Free ("pi", mk_permT (TFree ("'x", HOLogic.typeS))) $ - list_comb (c, args), - list_comb (c, map perm_arg (dts ~~ args))))) - end) constrs - end) descr; - - val (perm_simps, thy2) = - Primrec.add_primrec_overloaded - (map (fn (s, sT) => (s, sT, false)) - (List.take (perm_names' ~~ perm_names_types, length new_type_names))) - (map (fn s => (Binding.name s, NONE, NoSyn)) perm_names') perm_eqs thy1; - - (**** prove that permutation functions introduced by unfolding are ****) - (**** equivalent to already existing permutation functions ****) - - val _ = warning ("length descr: " ^ string_of_int (length descr)); - val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names)); - - val perm_indnames = DatatypeProp.make_tnames (map body_type perm_types); - val perm_fun_def = PureThy.get_thm thy2 "perm_fun_def"; - - val unfolded_perm_eq_thms = - if length descr = length new_type_names then [] - else map standard (List.drop (split_conj_thm - (Goal.prove_global thy2 [] [] - (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj - (map (fn (c as (s, T), x) => - let val [T1, T2] = binder_types T - in HOLogic.mk_eq (Const c $ pi $ Free (x, T2), - Const ("Nominal.perm", T) $ pi $ Free (x, T2)) - end) - (perm_names_types ~~ perm_indnames)))) - (fn _ => EVERY [indtac induction perm_indnames 1, - ALLGOALS (asm_full_simp_tac - (simpset_of thy2 addsimps [perm_fun_def]))])), - length new_type_names)); - - (**** prove [] \ t = t ****) - - val _ = warning "perm_empty_thms"; - - val perm_empty_thms = List.concat (map (fn a => - let val permT = mk_permT (Type (a, [])) - in map standard (List.take (split_conj_thm - (Goal.prove_global thy2 [] [] - (augment_sort thy2 [pt_class_of thy2 a] - (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj - (map (fn ((s, T), x) => HOLogic.mk_eq - (Const (s, permT --> T --> T) $ - Const ("List.list.Nil", permT) $ Free (x, T), - Free (x, T))) - (perm_names ~~ - map body_type perm_types ~~ perm_indnames))))) - (fn _ => EVERY [indtac induction perm_indnames 1, - ALLGOALS (asm_full_simp_tac (simpset_of thy2))])), - length new_type_names)) - end) - atoms); - - (**** prove (pi1 @ pi2) \ t = pi1 \ (pi2 \ t) ****) - - val _ = warning "perm_append_thms"; - - (*FIXME: these should be looked up statically*) - val at_pt_inst = PureThy.get_thm thy2 "at_pt_inst"; - val pt2 = PureThy.get_thm thy2 "pt2"; - - val perm_append_thms = List.concat (map (fn a => - let - val permT = mk_permT (Type (a, [])); - val pi1 = Free ("pi1", permT); - val pi2 = Free ("pi2", permT); - val pt_inst = pt_inst_of thy2 a; - val pt2' = pt_inst RS pt2; - val pt2_ax = PureThy.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "2") a); - in List.take (map standard (split_conj_thm - (Goal.prove_global thy2 [] [] - (augment_sort thy2 [pt_class_of thy2 a] - (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj - (map (fn ((s, T), x) => - let val perm = Const (s, permT --> T --> T) - in HOLogic.mk_eq - (perm $ (Const ("List.append", permT --> permT --> permT) $ - pi1 $ pi2) $ Free (x, T), - perm $ pi1 $ (perm $ pi2 $ Free (x, T))) - end) - (perm_names ~~ - map body_type perm_types ~~ perm_indnames))))) - (fn _ => EVERY [indtac induction perm_indnames 1, - ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt2', pt2_ax]))]))), - length new_type_names) - end) atoms); - - (**** prove pi1 ~ pi2 ==> pi1 \ t = pi2 \ t ****) - - val _ = warning "perm_eq_thms"; - - val pt3 = PureThy.get_thm thy2 "pt3"; - val pt3_rev = PureThy.get_thm thy2 "pt3_rev"; - - val perm_eq_thms = List.concat (map (fn a => - let - val permT = mk_permT (Type (a, [])); - val pi1 = Free ("pi1", permT); - val pi2 = Free ("pi2", permT); - val at_inst = at_inst_of thy2 a; - val pt_inst = pt_inst_of thy2 a; - val pt3' = pt_inst RS pt3; - val pt3_rev' = at_inst RS (pt_inst RS pt3_rev); - val pt3_ax = PureThy.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "3") a); - in List.take (map standard (split_conj_thm - (Goal.prove_global thy2 [] [] - (augment_sort thy2 [pt_class_of thy2 a] (Logic.mk_implies - (HOLogic.mk_Trueprop (Const ("Nominal.prm_eq", - permT --> permT --> HOLogic.boolT) $ pi1 $ pi2), - HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj - (map (fn ((s, T), x) => - let val perm = Const (s, permT --> T --> T) - in HOLogic.mk_eq - (perm $ pi1 $ Free (x, T), - perm $ pi2 $ Free (x, T)) - end) - (perm_names ~~ - map body_type perm_types ~~ perm_indnames)))))) - (fn _ => EVERY [indtac induction perm_indnames 1, - ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt3', pt3_rev', pt3_ax]))]))), - length new_type_names) - end) atoms); - - (**** prove pi1 \ (pi2 \ t) = (pi1 \ pi2) \ (pi1 \ t) ****) - - val cp1 = PureThy.get_thm thy2 "cp1"; - val dj_cp = PureThy.get_thm thy2 "dj_cp"; - val pt_perm_compose = PureThy.get_thm thy2 "pt_perm_compose"; - val pt_perm_compose_rev = PureThy.get_thm thy2 "pt_perm_compose_rev"; - val dj_perm_perm_forget = PureThy.get_thm thy2 "dj_perm_perm_forget"; - - fun composition_instance name1 name2 thy = - let - val cp_class = cp_class_of thy name1 name2; - val pt_class = - if name1 = name2 then [pt_class_of thy name1] - else []; - val permT1 = mk_permT (Type (name1, [])); - val permT2 = mk_permT (Type (name2, [])); - val Ts = map body_type perm_types; - val cp_inst = cp_inst_of thy name1 name2; - val simps = simpset_of thy addsimps (perm_fun_def :: - (if name1 <> name2 then - let val dj = dj_thm_of thy name2 name1 - in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end - else - let - val at_inst = at_inst_of thy name1; - val pt_inst = pt_inst_of thy name1; - in - [cp_inst RS cp1 RS sym, - at_inst RS (pt_inst RS pt_perm_compose) RS sym, - at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym] - end)) - val sort = Sign.certify_sort thy (cp_class :: pt_class); - val thms = split_conj_thm (Goal.prove_global thy [] [] - (augment_sort thy sort - (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj - (map (fn ((s, T), x) => - let - val pi1 = Free ("pi1", permT1); - val pi2 = Free ("pi2", permT2); - val perm1 = Const (s, permT1 --> T --> T); - val perm2 = Const (s, permT2 --> T --> T); - val perm3 = Const ("Nominal.perm", permT1 --> permT2 --> permT2) - in HOLogic.mk_eq - (perm1 $ pi1 $ (perm2 $ pi2 $ Free (x, T)), - perm2 $ (perm3 $ pi1 $ pi2) $ (perm1 $ pi1 $ Free (x, T))) - end) - (perm_names ~~ Ts ~~ perm_indnames))))) - (fn _ => EVERY [indtac induction perm_indnames 1, - ALLGOALS (asm_full_simp_tac simps)])) - in - fold (fn (s, tvs) => fn thy => AxClass.prove_arity - (s, map (inter_sort thy sort o snd) tvs, [cp_class]) - (Class.intro_classes_tac [] THEN ALLGOALS (resolve_tac thms)) thy) - (full_new_type_names' ~~ tyvars) thy - end; - - val (perm_thmss,thy3) = thy2 |> - fold (fn name1 => fold (composition_instance name1) atoms) atoms |> - fold (fn atom => fn thy => - let val pt_name = pt_class_of thy atom - in - fold (fn (s, tvs) => fn thy => AxClass.prove_arity - (s, map (inter_sort thy [pt_name] o snd) tvs, [pt_name]) - (EVERY - [Class.intro_classes_tac [], - resolve_tac perm_empty_thms 1, - resolve_tac perm_append_thms 1, - resolve_tac perm_eq_thms 1, assume_tac 1]) thy) - (full_new_type_names' ~~ tyvars) thy - end) atoms |> - PureThy.add_thmss - [((Binding.name (space_implode "_" new_type_names ^ "_unfolded_perm_eq"), - unfolded_perm_eq_thms), [Simplifier.simp_add]), - ((Binding.name (space_implode "_" new_type_names ^ "_perm_empty"), - perm_empty_thms), [Simplifier.simp_add]), - ((Binding.name (space_implode "_" new_type_names ^ "_perm_append"), - perm_append_thms), [Simplifier.simp_add]), - ((Binding.name (space_implode "_" new_type_names ^ "_perm_eq"), - perm_eq_thms), [Simplifier.simp_add])]; - - (**** Define representing sets ****) - - val _ = warning "representing sets"; - - val rep_set_names = DatatypeProp.indexify_names - (map (fn (i, _) => name_of_typ (nth_dtyp i) ^ "_set") descr); - val big_rep_name = - space_implode "_" (DatatypeProp.indexify_names (List.mapPartial - (fn (i, ("Nominal.noption", _, _)) => NONE - | (i, _) => SOME (name_of_typ (nth_dtyp i))) descr)) ^ "_set"; - val _ = warning ("big_rep_name: " ^ big_rep_name); - - fun strip_option (dtf as DtType ("fun", [dt, DtRec i])) = - (case AList.lookup op = descr i of - SOME ("Nominal.noption", _, [(_, [dt']), _]) => - apfst (cons dt) (strip_option dt') - | _ => ([], dtf)) - | strip_option (DtType ("fun", [dt, DtType ("Nominal.noption", [dt'])])) = - apfst (cons dt) (strip_option dt') - | strip_option dt = ([], dt); - - val dt_atomTs = distinct op = (map (typ_of_dtyp descr sorts) - (List.concat (map (fn (_, (_, _, cs)) => List.concat - (map (List.concat o map (fst o strip_option) o snd) cs)) descr))); - val dt_atoms = map (fst o dest_Type) dt_atomTs; - - fun make_intr s T (cname, cargs) = - let - fun mk_prem (dt, (j, j', prems, ts)) = - let - val (dts, dt') = strip_option dt; - val (dts', dt'') = strip_dtyp dt'; - val Ts = map (typ_of_dtyp descr sorts) dts; - val Us = map (typ_of_dtyp descr sorts) dts'; - val T = typ_of_dtyp descr sorts dt''; - val free = mk_Free "x" (Us ---> T) j; - val free' = app_bnds free (length Us); - fun mk_abs_fun (T, (i, t)) = - let val U = fastype_of t - in (i + 1, Const ("Nominal.abs_fun", [T, U, T] ---> - Type ("Nominal.noption", [U])) $ mk_Free "y" T i $ t) - end - in (j + 1, j' + length Ts, - case dt'' of - DtRec k => list_all (map (pair "x") Us, - HOLogic.mk_Trueprop (Free (List.nth (rep_set_names, k), - T --> HOLogic.boolT) $ free')) :: prems - | _ => prems, - snd (List.foldr mk_abs_fun (j', free) Ts) :: ts) - end; - - val (_, _, prems, ts) = List.foldr mk_prem (1, 1, [], []) cargs; - val concl = HOLogic.mk_Trueprop (Free (s, T --> HOLogic.boolT) $ - list_comb (Const (cname, map fastype_of ts ---> T), ts)) - in Logic.list_implies (prems, concl) - end; - - val (intr_ts, (rep_set_names', recTs')) = - apfst List.concat (apsnd ListPair.unzip (ListPair.unzip (List.mapPartial - (fn ((_, ("Nominal.noption", _, _)), _) => NONE - | ((i, (_, _, constrs)), rep_set_name) => - let val T = nth_dtyp i - in SOME (map (make_intr rep_set_name T) constrs, - (rep_set_name, T)) - end) - (descr ~~ rep_set_names)))); - val rep_set_names'' = map (Sign.full_bname thy3) rep_set_names'; - - val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy4) = - Inductive.add_inductive_global (serial_string ()) - {quiet_mode = false, verbose = false, kind = Thm.internalK, - alt_name = Binding.name big_rep_name, coind = false, no_elim = true, no_ind = false, - skip_mono = true, fork_mono = false} - (map (fn (s, T) => ((Binding.name s, T --> HOLogic.boolT), NoSyn)) - (rep_set_names' ~~ recTs')) - [] (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] thy3; - - (**** Prove that representing set is closed under permutation ****) - - val _ = warning "proving closure under permutation..."; - - val abs_perm = PureThy.get_thms thy4 "abs_perm"; - - val perm_indnames' = List.mapPartial - (fn (x, (_, ("Nominal.noption", _, _))) => NONE | (x, _) => SOME x) - (perm_indnames ~~ descr); - - fun mk_perm_closed name = map (fn th => standard (th RS mp)) - (List.take (split_conj_thm (Goal.prove_global thy4 [] [] - (augment_sort thy4 - (pt_class_of thy4 name :: map (cp_class_of thy4 name) (dt_atoms \ name)) - (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map - (fn ((s, T), x) => - let - val S = Const (s, T --> HOLogic.boolT); - val permT = mk_permT (Type (name, [])) - in HOLogic.mk_imp (S $ Free (x, T), - S $ (Const ("Nominal.perm", permT --> T --> T) $ - Free ("pi", permT) $ Free (x, T))) - end) (rep_set_names'' ~~ recTs' ~~ perm_indnames'))))) - (fn _ => EVERY - [indtac rep_induct [] 1, - ALLGOALS (simp_tac (simpset_of thy4 addsimps - (symmetric perm_fun_def :: abs_perm))), - ALLGOALS (resolve_tac rep_intrs THEN_ALL_NEW assume_tac)])), - length new_type_names)); - - val perm_closed_thmss = map mk_perm_closed atoms; - - (**** typedef ****) - - val _ = warning "defining type..."; - - val (typedefs, thy6) = - thy4 - |> fold_map (fn ((((name, mx), tvs), (cname, U)), name') => fn thy => - Typedef.add_typedef false (SOME (Binding.name name')) - (Binding.name name, map fst tvs, mx) - (Const ("Collect", (U --> HOLogic.boolT) --> HOLogic.mk_setT U) $ - Const (cname, U --> HOLogic.boolT)) NONE - (rtac exI 1 THEN rtac CollectI 1 THEN - QUIET_BREADTH_FIRST (has_fewer_prems 1) - (resolve_tac rep_intrs 1)) thy |> (fn ((_, r), thy) => - let - val permT = mk_permT - (TFree (Name.variant (map fst tvs) "'a", HOLogic.typeS)); - val pi = Free ("pi", permT); - val T = Type (Sign.intern_type thy name, map TFree tvs); - in apfst (pair r o hd) - (PureThy.add_defs_unchecked true [((Binding.name ("prm_" ^ name ^ "_def"), Logic.mk_equals - (Const ("Nominal.perm", permT --> T --> T) $ pi $ Free ("x", T), - Const (Sign.intern_const thy ("Abs_" ^ name), U --> T) $ - (Const ("Nominal.perm", permT --> U --> U) $ pi $ - (Const (Sign.intern_const thy ("Rep_" ^ name), T --> U) $ - Free ("x", T))))), [])] thy) - end)) - (types_syntax ~~ tyvars ~~ - List.take (rep_set_names'' ~~ recTs', length new_type_names) ~~ - new_type_names); - - val perm_defs = map snd typedefs; - val Abs_inverse_thms = map (collect_simp o #Abs_inverse o fst) typedefs; - val Rep_inverse_thms = map (#Rep_inverse o fst) typedefs; - val Rep_thms = map (collect_simp o #Rep o fst) typedefs; - - - (** prove that new types are in class pt_ **) - - val _ = warning "prove that new types are in class pt_ ..."; - - fun pt_instance (atom, perm_closed_thms) = - fold (fn ((((((Abs_inverse, Rep_inverse), Rep), - perm_def), name), tvs), perm_closed) => fn thy => - let - val pt_class = pt_class_of thy atom; - val sort = Sign.certify_sort thy - (pt_class :: map (cp_class_of thy atom) (dt_atoms \ atom)) - in AxClass.prove_arity - (Sign.intern_type thy name, - map (inter_sort thy sort o snd) tvs, [pt_class]) - (EVERY [Class.intro_classes_tac [], - rewrite_goals_tac [perm_def], - asm_full_simp_tac (simpset_of thy addsimps [Rep_inverse]) 1, - asm_full_simp_tac (simpset_of thy addsimps - [Rep RS perm_closed RS Abs_inverse]) 1, - asm_full_simp_tac (HOL_basic_ss addsimps [PureThy.get_thm thy - ("pt_" ^ Long_Name.base_name atom ^ "3")]) 1]) thy - end) - (Abs_inverse_thms ~~ Rep_inverse_thms ~~ Rep_thms ~~ perm_defs ~~ - new_type_names ~~ tyvars ~~ perm_closed_thms); - - - (** prove that new types are in class cp__ **) - - val _ = warning "prove that new types are in class cp__ ..."; - - fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy = - let - val cp_class = cp_class_of thy atom1 atom2; - val sort = Sign.certify_sort thy - (pt_class_of thy atom1 :: map (cp_class_of thy atom1) (dt_atoms \ atom1) @ - (if atom1 = atom2 then [cp_class_of thy atom1 atom1] else - pt_class_of thy atom2 :: map (cp_class_of thy atom2) (dt_atoms \ atom2))); - val cp1' = cp_inst_of thy atom1 atom2 RS cp1 - in fold (fn ((((((Abs_inverse, Rep), - perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy => - AxClass.prove_arity - (Sign.intern_type thy name, - map (inter_sort thy sort o snd) tvs, [cp_class]) - (EVERY [Class.intro_classes_tac [], - rewrite_goals_tac [perm_def], - asm_full_simp_tac (simpset_of thy addsimps - ((Rep RS perm_closed1 RS Abs_inverse) :: - (if atom1 = atom2 then [] - else [Rep RS perm_closed2 RS Abs_inverse]))) 1, - cong_tac 1, - rtac refl 1, - rtac cp1' 1]) thy) - (Abs_inverse_thms ~~ Rep_thms ~~ perm_defs ~~ new_type_names ~~ - tyvars ~~ perm_closed_thms1 ~~ perm_closed_thms2) thy - end; - - val thy7 = fold (fn x => fn thy => thy |> - pt_instance x |> - fold (cp_instance x) (atoms ~~ perm_closed_thmss)) - (atoms ~~ perm_closed_thmss) thy6; - - (**** constructors ****) - - fun mk_abs_fun (x, t) = - let - val T = fastype_of x; - val U = fastype_of t - in - Const ("Nominal.abs_fun", T --> U --> T --> - Type ("Nominal.noption", [U])) $ x $ t - end; - - val (ty_idxs, _) = List.foldl - (fn ((i, ("Nominal.noption", _, _)), p) => p - | ((i, _), (ty_idxs, j)) => (ty_idxs @ [(i, j)], j + 1)) ([], 0) descr; - - fun reindex (DtType (s, dts)) = DtType (s, map reindex dts) - | reindex (DtRec i) = DtRec (the (AList.lookup op = ty_idxs i)) - | reindex dt = dt; - - fun strip_suffix i s = implode (List.take (explode s, size s - i)); - - (** strips the "_Rep" in type names *) - fun strip_nth_name i s = - let val xs = Long_Name.explode s; - in Long_Name.implode (Library.nth_map (length xs - i) (strip_suffix 4) xs) end; - - val (descr'', ndescr) = ListPair.unzip (map_filter - (fn (i, ("Nominal.noption", _, _)) => NONE - | (i, (s, dts, constrs)) => - let - val SOME index = AList.lookup op = ty_idxs i; - val (constrs2, constrs1) = - map_split (fn (cname, cargs) => - apsnd (pair (strip_nth_name 2 (strip_nth_name 1 cname))) - (fold_map (fn dt => fn dts => - let val (dts', dt') = strip_option dt - in ((length dts, length dts'), dts @ dts' @ [reindex dt']) end) - cargs [])) constrs - in SOME ((index, (strip_nth_name 1 s, map reindex dts, constrs1)), - (index, constrs2)) - end) descr); - - val (descr1, descr2) = chop (length new_type_names) descr''; - val descr' = [descr1, descr2]; - - fun partition_cargs idxs xs = map (fn (i, j) => - (List.take (List.drop (xs, i), j), List.nth (xs, i + j))) idxs; - - val pdescr = map (fn ((i, (s, dts, constrs)), (_, idxss)) => (i, (s, dts, - map (fn ((cname, cargs), idxs) => (cname, partition_cargs idxs cargs)) - (constrs ~~ idxss)))) (descr'' ~~ ndescr); - - fun nth_dtyp' i = typ_of_dtyp descr'' sorts (DtRec i); - - val rep_names = map (fn s => - Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names; - val abs_names = map (fn s => - Sign.intern_const thy7 ("Abs_" ^ s)) new_type_names; - - val recTs = get_rec_types descr'' sorts; - val newTs' = Library.take (length new_type_names, recTs'); - val newTs = Library.take (length new_type_names, recTs); - - val full_new_type_names = map (Sign.full_bname thy) new_type_names; - - fun make_constr_def tname T T' ((thy, defs, eqns), - (((cname_rep, _), (cname, cargs)), (cname', mx))) = - let - fun constr_arg ((dts, dt), (j, l_args, r_args)) = - let - val xs = map (fn (dt, i) => mk_Free "x" (typ_of_dtyp descr'' sorts dt) i) - (dts ~~ (j upto j + length dts - 1)) - val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts) - in - (j + length dts + 1, - xs @ x :: l_args, - List.foldr mk_abs_fun - (case dt of - DtRec k => if k < length new_type_names then - Const (List.nth (rep_names, k), typ_of_dtyp descr'' sorts dt --> - typ_of_dtyp descr sorts dt) $ x - else error "nested recursion not (yet) supported" - | _ => x) xs :: r_args) - end - - val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) cargs; - val abs_name = Sign.intern_const thy ("Abs_" ^ tname); - val rep_name = Sign.intern_const thy ("Rep_" ^ tname); - val constrT = map fastype_of l_args ---> T; - val lhs = list_comb (Const (cname, constrT), l_args); - val rhs = list_comb (Const (cname_rep, map fastype_of r_args ---> T'), r_args); - val def = Logic.mk_equals (lhs, Const (abs_name, T' --> T) $ rhs); - val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq - (Const (rep_name, T --> T') $ lhs, rhs)); - val def_name = (Long_Name.base_name cname) ^ "_def"; - val ([def_thm], thy') = thy |> - Sign.add_consts_i [(Binding.name cname', constrT, mx)] |> - (PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)] - in (thy', defs @ [def_thm], eqns @ [eqn]) end; - - fun dt_constr_defs ((thy, defs, eqns, dist_lemmas), ((((((_, (_, _, constrs)), - (_, (_, _, constrs'))), tname), T), T'), constr_syntax)) = - let - val rep_const = cterm_of thy - (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T')); - val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma); - val (thy', defs', eqns') = Library.foldl (make_constr_def tname T T') - ((Sign.add_path tname thy, defs, []), constrs ~~ constrs' ~~ constr_syntax) - in - (parent_path (#flat_names config) thy', defs', eqns @ [eqns'], dist_lemmas @ [dist]) - end; - - val (thy8, constr_defs, constr_rep_eqns, dist_lemmas) = Library.foldl dt_constr_defs - ((thy7, [], [], []), List.take (descr, length new_type_names) ~~ - List.take (pdescr, length new_type_names) ~~ - new_type_names ~~ newTs ~~ newTs' ~~ constr_syntax); - - val abs_inject_thms = map (collect_simp o #Abs_inject o fst) typedefs - val rep_inject_thms = map (#Rep_inject o fst) typedefs - - (* prove theorem Rep_i (Constr_j ...) = Constr'_j ... *) - - fun prove_constr_rep_thm eqn = - let - val inj_thms = map (fn r => r RS iffD1) abs_inject_thms; - val rewrites = constr_defs @ map mk_meta_eq Rep_inverse_thms - in Goal.prove_global thy8 [] [] eqn (fn _ => EVERY - [resolve_tac inj_thms 1, - rewrite_goals_tac rewrites, - rtac refl 3, - resolve_tac rep_intrs 2, - REPEAT (resolve_tac Rep_thms 1)]) - end; - - val constr_rep_thmss = map (map prove_constr_rep_thm) constr_rep_eqns; - - (* prove theorem pi \ Rep_i x = Rep_i (pi \ x) *) - - fun prove_perm_rep_perm (atom, perm_closed_thms) = map (fn th => - let - val _ $ (_ $ (Rep $ x)) = Logic.unvarify (prop_of th); - val Type ("fun", [T, U]) = fastype_of Rep; - val permT = mk_permT (Type (atom, [])); - val pi = Free ("pi", permT); - in - Goal.prove_global thy8 [] [] - (augment_sort thy8 - (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (dt_atoms \ atom)) - (HOLogic.mk_Trueprop (HOLogic.mk_eq - (Const ("Nominal.perm", permT --> U --> U) $ pi $ (Rep $ x), - Rep $ (Const ("Nominal.perm", permT --> T --> T) $ pi $ x))))) - (fn _ => simp_tac (HOL_basic_ss addsimps (perm_defs @ Abs_inverse_thms @ - perm_closed_thms @ Rep_thms)) 1) - end) Rep_thms; - - val perm_rep_perm_thms = List.concat (map prove_perm_rep_perm - (atoms ~~ perm_closed_thmss)); - - (* prove distinctness theorems *) - - val distinct_props = DatatypeProp.make_distincts descr' sorts; - val dist_rewrites = map2 (fn rep_thms => fn dist_lemma => - dist_lemma :: rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0]) - constr_rep_thmss dist_lemmas; - - fun prove_distinct_thms _ (_, []) = [] - | prove_distinct_thms (p as (rep_thms, dist_lemma)) (k, t :: ts) = - let - val dist_thm = Goal.prove_global thy8 [] [] t (fn _ => - simp_tac (simpset_of thy8 addsimps (dist_lemma :: rep_thms)) 1) - in dist_thm :: standard (dist_thm RS not_sym) :: - prove_distinct_thms p (k, ts) - end; - - val distinct_thms = map2 prove_distinct_thms - (constr_rep_thmss ~~ dist_lemmas) distinct_props; - - (** prove equations for permutation functions **) - - val perm_simps' = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) => - let val T = nth_dtyp' i - in List.concat (map (fn (atom, perm_closed_thms) => - map (fn ((cname, dts), constr_rep_thm) => - let - val cname = Sign.intern_const thy8 - (Long_Name.append tname (Long_Name.base_name cname)); - val permT = mk_permT (Type (atom, [])); - val pi = Free ("pi", permT); - - fun perm t = - let val T = fastype_of t - in Const ("Nominal.perm", permT --> T --> T) $ pi $ t end; - - fun constr_arg ((dts, dt), (j, l_args, r_args)) = - let - val Ts = map (typ_of_dtyp descr'' sorts) dts; - val xs = map (fn (T, i) => mk_Free "x" T i) - (Ts ~~ (j upto j + length dts - 1)) - val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts) - in - (j + length dts + 1, - xs @ x :: l_args, - map perm (xs @ [x]) @ r_args) - end - - val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) dts; - val c = Const (cname, map fastype_of l_args ---> T) - in - Goal.prove_global thy8 [] [] - (augment_sort thy8 - (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (dt_atoms \ atom)) - (HOLogic.mk_Trueprop (HOLogic.mk_eq - (perm (list_comb (c, l_args)), list_comb (c, r_args))))) - (fn _ => EVERY - [simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: perm_defs)) 1, - simp_tac (HOL_basic_ss addsimps (Rep_thms @ Abs_inverse_thms @ - constr_defs @ perm_closed_thms)) 1, - TRY (simp_tac (HOL_basic_ss addsimps - (symmetric perm_fun_def :: abs_perm)) 1), - TRY (simp_tac (HOL_basic_ss addsimps - (perm_fun_def :: perm_defs @ Rep_thms @ Abs_inverse_thms @ - perm_closed_thms)) 1)]) - end) (constrs ~~ constr_rep_thms)) (atoms ~~ perm_closed_thmss)) - end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss); - - (** prove injectivity of constructors **) - - val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms; - val alpha = PureThy.get_thms thy8 "alpha"; - val abs_fresh = PureThy.get_thms thy8 "abs_fresh"; - - val pt_cp_sort = - map (pt_class_of thy8) dt_atoms @ - maps (fn s => map (cp_class_of thy8 s) (dt_atoms \ s)) dt_atoms; - - val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) => - let val T = nth_dtyp' i - in List.mapPartial (fn ((cname, dts), constr_rep_thm) => - if null dts then NONE else SOME - let - val cname = Sign.intern_const thy8 - (Long_Name.append tname (Long_Name.base_name cname)); - - fun make_inj ((dts, dt), (j, args1, args2, eqs)) = - let - val Ts_idx = map (typ_of_dtyp descr'' sorts) dts ~~ (j upto j + length dts - 1); - val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx; - val ys = map (fn (T, i) => mk_Free "y" T i) Ts_idx; - val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts); - val y = mk_Free "y" (typ_of_dtyp descr'' sorts dt) (j + length dts) - in - (j + length dts + 1, - xs @ (x :: args1), ys @ (y :: args2), - HOLogic.mk_eq - (List.foldr mk_abs_fun x xs, List.foldr mk_abs_fun y ys) :: eqs) - end; - - val (_, args1, args2, eqs) = List.foldr make_inj (1, [], [], []) dts; - val Ts = map fastype_of args1; - val c = Const (cname, Ts ---> T) - in - Goal.prove_global thy8 [] [] - (augment_sort thy8 pt_cp_sort - (HOLogic.mk_Trueprop (HOLogic.mk_eq - (HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)), - foldr1 HOLogic.mk_conj eqs)))) - (fn _ => EVERY - [asm_full_simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: - rep_inject_thms')) 1, - TRY (asm_full_simp_tac (HOL_basic_ss addsimps (fresh_def :: supp_def :: - alpha @ abs_perm @ abs_fresh @ rep_inject_thms @ - perm_rep_perm_thms)) 1)]) - end) (constrs ~~ constr_rep_thms) - end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss); - - (** equations for support and freshness **) - - val (supp_thms, fresh_thms) = ListPair.unzip (map ListPair.unzip - (map (fn ((((i, (_, _, constrs)), tname), inject_thms'), perm_thms') => - let val T = nth_dtyp' i - in List.concat (map (fn (cname, dts) => map (fn atom => - let - val cname = Sign.intern_const thy8 - (Long_Name.append tname (Long_Name.base_name cname)); - val atomT = Type (atom, []); - - fun process_constr ((dts, dt), (j, args1, args2)) = - let - val Ts_idx = map (typ_of_dtyp descr'' sorts) dts ~~ (j upto j + length dts - 1); - val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx; - val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts) - in - (j + length dts + 1, - xs @ (x :: args1), List.foldr mk_abs_fun x xs :: args2) - end; - - val (_, args1, args2) = List.foldr process_constr (1, [], []) dts; - val Ts = map fastype_of args1; - val c = list_comb (Const (cname, Ts ---> T), args1); - fun supp t = - Const ("Nominal.supp", fastype_of t --> HOLogic.mk_setT atomT) $ t; - fun fresh t = fresh_const atomT (fastype_of t) $ Free ("a", atomT) $ t; - val supp_thm = Goal.prove_global thy8 [] [] - (augment_sort thy8 pt_cp_sort - (HOLogic.mk_Trueprop (HOLogic.mk_eq - (supp c, - if null dts then HOLogic.mk_set atomT [] - else foldr1 (HOLogic.mk_binop @{const_name Un}) (map supp args2))))) - (fn _ => - simp_tac (HOL_basic_ss addsimps (supp_def :: - Un_assoc :: de_Morgan_conj :: Collect_disj_eq :: finite_Un :: - symmetric empty_def :: finite_emptyI :: simp_thms @ - abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1) - in - (supp_thm, - Goal.prove_global thy8 [] [] (augment_sort thy8 pt_cp_sort - (HOLogic.mk_Trueprop (HOLogic.mk_eq - (fresh c, - if null dts then HOLogic.true_const - else foldr1 HOLogic.mk_conj (map fresh args2))))) - (fn _ => - simp_tac (HOL_ss addsimps [Un_iff, empty_iff, fresh_def, supp_thm]) 1)) - end) atoms) constrs) - end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ inject_thms ~~ perm_simps'))); - - (**** weak induction theorem ****) - - fun mk_indrule_lemma ((prems, concls), (((i, _), T), U)) = - let - val Rep_t = Const (List.nth (rep_names, i), T --> U) $ - mk_Free "x" T i; - - val Abs_t = Const (List.nth (abs_names, i), U --> T) - - in (prems @ [HOLogic.imp $ - (Const (List.nth (rep_set_names'', i), U --> HOLogic.boolT) $ Rep_t) $ - (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))], - concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i]) - end; - - val (indrule_lemma_prems, indrule_lemma_concls) = - Library.foldl mk_indrule_lemma (([], []), (descr'' ~~ recTs ~~ recTs')); - - val indrule_lemma = Goal.prove_global thy8 [] [] - (Logic.mk_implies - (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems), - HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY - [REPEAT (etac conjE 1), - REPEAT (EVERY - [TRY (rtac conjI 1), full_simp_tac (HOL_basic_ss addsimps Rep_inverse_thms) 1, - etac mp 1, resolve_tac Rep_thms 1])]); - - val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma))); - val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else - map (Free o apfst fst o dest_Var) Ps; - val indrule_lemma' = cterm_instantiate - (map (cterm_of thy8) Ps ~~ map (cterm_of thy8) frees) indrule_lemma; - - val Abs_inverse_thms' = map (fn r => r RS subst) Abs_inverse_thms; - - val dt_induct_prop = DatatypeProp.make_ind descr' sorts; - val dt_induct = Goal.prove_global thy8 [] - (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop) - (fn {prems, ...} => EVERY - [rtac indrule_lemma' 1, - (indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, - EVERY (map (fn (prem, r) => (EVERY - [REPEAT (eresolve_tac Abs_inverse_thms' 1), - simp_tac (HOL_basic_ss addsimps [symmetric r]) 1, - DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)])) - (prems ~~ constr_defs))]); - - val case_names_induct = mk_case_names_induct descr''; - - (**** prove that new datatypes have finite support ****) - - val _ = warning "proving finite support for the new datatype"; - - val indnames = DatatypeProp.make_tnames recTs; - - val abs_supp = PureThy.get_thms thy8 "abs_supp"; - val supp_atm = PureThy.get_thms thy8 "supp_atm"; - - val finite_supp_thms = map (fn atom => - let val atomT = Type (atom, []) - in map standard (List.take - (split_conj_thm (Goal.prove_global thy8 [] [] - (augment_sort thy8 (fs_class_of thy8 atom :: pt_cp_sort) - (HOLogic.mk_Trueprop - (foldr1 HOLogic.mk_conj (map (fn (s, T) => - Const ("Finite_Set.finite", HOLogic.mk_setT atomT --> HOLogic.boolT) $ - (Const ("Nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T))) - (indnames ~~ recTs))))) - (fn _ => indtac dt_induct indnames 1 THEN - ALLGOALS (asm_full_simp_tac (simpset_of thy8 addsimps - (abs_supp @ supp_atm @ - PureThy.get_thms thy8 ("fs_" ^ Long_Name.base_name atom ^ "1") @ - List.concat supp_thms))))), - length new_type_names)) - end) atoms; - - val simp_atts = replicate (length new_type_names) [Simplifier.simp_add]; - - (* Function to add both the simp and eqvt attributes *) - (* These two attributes are duplicated on all the types in the mutual nominal datatypes *) - - val simp_eqvt_atts = replicate (length new_type_names) [Simplifier.simp_add, NominalThmDecls.eqvt_add]; - - val (_, thy9) = thy8 |> - Sign.add_path big_name |> - PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])] ||>> - PureThy.add_thmss [((Binding.name "inducts", projections dt_induct), [case_names_induct])] ||> - Sign.parent_path ||>> - DatatypeAux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>> - DatatypeAux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>> - DatatypeAux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>> - DatatypeAux.store_thmss "inject" new_type_names inject_thms ||>> - DatatypeAux.store_thmss "supp" new_type_names supp_thms ||>> - DatatypeAux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||> - fold (fn (atom, ths) => fn thy => - let - val class = fs_class_of thy atom; - val sort = Sign.certify_sort thy (class :: pt_cp_sort) - in fold (fn Type (s, Ts) => AxClass.prove_arity - (s, map (inter_sort thy sort o snd o dest_TFree) Ts, [class]) - (Class.intro_classes_tac [] THEN resolve_tac ths 1)) newTs thy - end) (atoms ~~ finite_supp_thms); - - (**** strong induction theorem ****) - - val pnames = if length descr'' = 1 then ["P"] - else map (fn i => "P" ^ string_of_int i) (1 upto length descr''); - val ind_sort = if null dt_atomTs then HOLogic.typeS - else Sign.certify_sort thy9 (map (fs_class_of thy9) dt_atoms); - val fsT = TFree ("'n", ind_sort); - val fsT' = TFree ("'n", HOLogic.typeS); - - val fresh_fs = map (fn (s, T) => (T, Free (s, fsT' --> HOLogic.mk_setT T))) - (DatatypeProp.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs); - - fun make_pred fsT i T = - Free (List.nth (pnames, i), fsT --> T --> HOLogic.boolT); - - fun mk_fresh1 xs [] = [] - | mk_fresh1 xs ((y as (_, T)) :: ys) = map (fn x => HOLogic.mk_Trueprop - (HOLogic.mk_not (HOLogic.mk_eq (Free y, Free x)))) - (filter (fn (_, U) => T = U) (rev xs)) @ - mk_fresh1 (y :: xs) ys; - - fun mk_fresh2 xss [] = [] - | mk_fresh2 xss ((p as (ys, _)) :: yss) = List.concat (map (fn y as (_, T) => - map (fn (_, x as (_, U)) => HOLogic.mk_Trueprop - (fresh_const T U $ Free y $ Free x)) (rev xss @ yss)) ys) @ - mk_fresh2 (p :: xss) yss; - - fun make_ind_prem fsT f k T ((cname, cargs), idxs) = - let - val recs = List.filter is_rec_type cargs; - val Ts = map (typ_of_dtyp descr'' sorts) cargs; - val recTs' = map (typ_of_dtyp descr'' sorts) recs; - val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts); - val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs)); - val frees = tnames ~~ Ts; - val frees' = partition_cargs idxs frees; - val z = (Name.variant tnames "z", fsT); - - fun mk_prem ((dt, s), T) = - let - val (Us, U) = strip_type T; - val l = length Us - in list_all (z :: map (pair "x") Us, HOLogic.mk_Trueprop - (make_pred fsT (body_index dt) U $ Bound l $ app_bnds (Free (s, T)) l)) - end; - - val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs'); - val prems' = map (fn p as (_, T) => HOLogic.mk_Trueprop - (f T (Free p) (Free z))) (List.concat (map fst frees')) @ - mk_fresh1 [] (List.concat (map fst frees')) @ - mk_fresh2 [] frees' - - in list_all_free (frees @ [z], Logic.list_implies (prems' @ prems, - HOLogic.mk_Trueprop (make_pred fsT k T $ Free z $ - list_comb (Const (cname, Ts ---> T), map Free frees)))) - end; - - val ind_prems = List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) => - map (make_ind_prem fsT (fn T => fn t => fn u => - fresh_const T fsT $ t $ u) i T) - (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs)); - val tnames = DatatypeProp.make_tnames recTs; - val zs = Name.variant_list tnames (replicate (length descr'') "z"); - val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &") - (map (fn ((((i, _), T), tname), z) => - make_pred fsT i T $ Free (z, fsT) $ Free (tname, T)) - (descr'' ~~ recTs ~~ tnames ~~ zs))); - val induct = Logic.list_implies (ind_prems, ind_concl); - - val ind_prems' = - map (fn (_, f as Free (_, T)) => list_all_free ([("x", fsT')], - HOLogic.mk_Trueprop (Const ("Finite_Set.finite", - (snd (split_last (binder_types T)) --> HOLogic.boolT) --> - HOLogic.boolT) $ (f $ Free ("x", fsT'))))) fresh_fs @ - List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) => - map (make_ind_prem fsT' (fn T => fn t => fn u => HOLogic.Not $ - HOLogic.mk_mem (t, the (AList.lookup op = fresh_fs T) $ u)) i T) - (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs)); - val ind_concl' = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &") - (map (fn ((((i, _), T), tname), z) => - make_pred fsT' i T $ Free (z, fsT') $ Free (tname, T)) - (descr'' ~~ recTs ~~ tnames ~~ zs))); - val induct' = Logic.list_implies (ind_prems', ind_concl'); - - val aux_ind_vars = - (DatatypeProp.indexify_names (replicate (length dt_atomTs) "pi") ~~ - map mk_permT dt_atomTs) @ [("z", fsT')]; - val aux_ind_Ts = rev (map snd aux_ind_vars); - val aux_ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &") - (map (fn (((i, _), T), tname) => - HOLogic.list_all (aux_ind_vars, make_pred fsT' i T $ Bound 0 $ - fold_rev (mk_perm aux_ind_Ts) (map Bound (length dt_atomTs downto 1)) - (Free (tname, T)))) - (descr'' ~~ recTs ~~ tnames))); - - val fin_set_supp = map (fn s => - at_inst_of thy9 s RS at_fin_set_supp) dt_atoms; - val fin_set_fresh = map (fn s => - at_inst_of thy9 s RS at_fin_set_fresh) dt_atoms; - val pt1_atoms = map (fn Type (s, _) => - PureThy.get_thm thy9 ("pt_" ^ Long_Name.base_name s ^ "1")) dt_atomTs; - val pt2_atoms = map (fn Type (s, _) => - PureThy.get_thm thy9 ("pt_" ^ Long_Name.base_name s ^ "2") RS sym) dt_atomTs; - val exists_fresh' = PureThy.get_thms thy9 "exists_fresh'"; - val fs_atoms = PureThy.get_thms thy9 "fin_supp"; - val abs_supp = PureThy.get_thms thy9 "abs_supp"; - val perm_fresh_fresh = PureThy.get_thms thy9 "perm_fresh_fresh"; - val calc_atm = PureThy.get_thms thy9 "calc_atm"; - val fresh_atm = PureThy.get_thms thy9 "fresh_atm"; - val fresh_left = PureThy.get_thms thy9 "fresh_left"; - val perm_swap = PureThy.get_thms thy9 "perm_swap"; - - fun obtain_fresh_name' ths ts T (freshs1, freshs2, ctxt) = - let - val p = foldr1 HOLogic.mk_prod (ts @ freshs1); - val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop - (HOLogic.exists_const T $ Abs ("x", T, - fresh_const T (fastype_of p) $ - Bound 0 $ p))) - (fn _ => EVERY - [resolve_tac exists_fresh' 1, - simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms @ - fin_set_supp @ ths)) 1]); - val (([cx], ths), ctxt') = Obtain.result - (fn _ => EVERY - [etac exE 1, - full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1, - REPEAT (etac conjE 1)]) - [ex] ctxt - in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end; - - fun fresh_fresh_inst thy a b = - let - val T = fastype_of a; - val SOME th = find_first (fn th => case prop_of th of - _ $ (_ $ (Const (_, Type (_, [U, _])) $ _ $ _)) $ _ => U = T - | _ => false) perm_fresh_fresh - in - Drule.instantiate' [] - [SOME (cterm_of thy a), NONE, SOME (cterm_of thy b)] th - end; - - val fs_cp_sort = - map (fs_class_of thy9) dt_atoms @ - maps (fn s => map (cp_class_of thy9 s) (dt_atoms \ s)) dt_atoms; - - (********************************************************************** - The subgoals occurring in the proof of induct_aux have the - following parameters: - - x_1 ... x_k p_1 ... p_m z - - where - - x_i : constructor arguments (introduced by weak induction rule) - p_i : permutations (one for each atom type in the data type) - z : freshness context - ***********************************************************************) - - val _ = warning "proving strong induction theorem ..."; - - val induct_aux = Goal.prove_global thy9 [] - (map (augment_sort thy9 fs_cp_sort) ind_prems') - (augment_sort thy9 fs_cp_sort ind_concl') (fn {prems, context} => - let - val (prems1, prems2) = chop (length dt_atomTs) prems; - val ind_ss2 = HOL_ss addsimps - finite_Diff :: abs_fresh @ abs_supp @ fs_atoms; - val ind_ss1 = ind_ss2 addsimps fresh_left @ calc_atm @ - fresh_atm @ rev_simps @ app_simps; - val ind_ss3 = HOL_ss addsimps abs_fun_eq1 :: - abs_perm @ calc_atm @ perm_swap; - val ind_ss4 = HOL_basic_ss addsimps fresh_left @ prems1 @ - fin_set_fresh @ calc_atm; - val ind_ss5 = HOL_basic_ss addsimps pt1_atoms; - val ind_ss6 = HOL_basic_ss addsimps flat perm_simps'; - val th = Goal.prove context [] [] - (augment_sort thy9 fs_cp_sort aux_ind_concl) - (fn {context = context1, ...} => - EVERY (indtac dt_induct tnames 1 :: - maps (fn ((_, (_, _, constrs)), (_, constrs')) => - map (fn ((cname, cargs), is) => - REPEAT (rtac allI 1) THEN - SUBPROOF (fn {prems = iprems, params, concl, - context = context2, ...} => - let - val concl' = term_of concl; - val _ $ (_ $ _ $ u) = concl'; - val U = fastype_of u; - val (xs, params') = - chop (length cargs) (map term_of params); - val Ts = map fastype_of xs; - val cnstr = Const (cname, Ts ---> U); - val (pis, z) = split_last params'; - val mk_pi = fold_rev (mk_perm []) pis; - val xs' = partition_cargs is xs; - val xs'' = map (fn (ts, u) => (map mk_pi ts, mk_pi u)) xs'; - val ts = maps (fn (ts, u) => ts @ [u]) xs''; - val (freshs1, freshs2, context3) = fold (fn t => - let val T = fastype_of t - in obtain_fresh_name' prems1 - (the (AList.lookup op = fresh_fs T) $ z :: ts) T - end) (maps fst xs') ([], [], context2); - val freshs1' = unflat (map fst xs') freshs1; - val freshs2' = map (Simplifier.simplify ind_ss4) - (mk_not_sym freshs2); - val ind_ss1' = ind_ss1 addsimps freshs2'; - val ind_ss3' = ind_ss3 addsimps freshs2'; - val rename_eq = - if forall (null o fst) xs' then [] - else [Goal.prove context3 [] [] - (HOLogic.mk_Trueprop (HOLogic.mk_eq - (list_comb (cnstr, ts), - list_comb (cnstr, maps (fn ((bs, t), cs) => - cs @ [fold_rev (mk_perm []) (map perm_of_pair - (bs ~~ cs)) t]) (xs'' ~~ freshs1'))))) - (fn _ => EVERY - (simp_tac (HOL_ss addsimps flat inject_thms) 1 :: - REPEAT (FIRSTGOAL (rtac conjI)) :: - maps (fn ((bs, t), cs) => - if null bs then [] - else rtac sym 1 :: maps (fn (b, c) => - [rtac trans 1, rtac sym 1, - rtac (fresh_fresh_inst thy9 b c) 1, - simp_tac ind_ss1' 1, - simp_tac ind_ss2 1, - simp_tac ind_ss3' 1]) (bs ~~ cs)) - (xs'' ~~ freshs1')))]; - val th = Goal.prove context3 [] [] concl' (fn _ => EVERY - [simp_tac (ind_ss6 addsimps rename_eq) 1, - cut_facts_tac iprems 1, - (resolve_tac prems THEN_ALL_NEW - SUBGOAL (fn (t, i) => case Logic.strip_assums_concl t of - _ $ (Const ("Nominal.fresh", _) $ _ $ _) => - simp_tac ind_ss1' i - | _ $ (Const ("Not", _) $ _) => - resolve_tac freshs2' i - | _ => asm_simp_tac (HOL_basic_ss addsimps - pt2_atoms addsimprocs [perm_simproc]) i)) 1]) - val final = ProofContext.export context3 context2 [th] - in - resolve_tac final 1 - end) context1 1) (constrs ~~ constrs')) (descr'' ~~ ndescr))) - in - EVERY - [cut_facts_tac [th] 1, - REPEAT (eresolve_tac [conjE, @{thm allE_Nil}] 1), - REPEAT (etac allE 1), - REPEAT (TRY (rtac conjI 1) THEN asm_full_simp_tac ind_ss5 1)] - end); - - val induct_aux' = Thm.instantiate ([], - map (fn (s, v as Var (_, T)) => - (cterm_of thy9 v, cterm_of thy9 (Free (s, T)))) - (pnames ~~ map head_of (HOLogic.dest_conj - (HOLogic.dest_Trueprop (concl_of induct_aux)))) @ - map (fn (_, f) => - let val f' = Logic.varify f - in (cterm_of thy9 f', - cterm_of thy9 (Const ("Nominal.supp", fastype_of f'))) - end) fresh_fs) induct_aux; - - val induct = Goal.prove_global thy9 [] - (map (augment_sort thy9 fs_cp_sort) ind_prems) - (augment_sort thy9 fs_cp_sort ind_concl) - (fn {prems, ...} => EVERY - [rtac induct_aux' 1, - REPEAT (resolve_tac fs_atoms 1), - REPEAT ((resolve_tac prems THEN_ALL_NEW - (etac meta_spec ORELSE' full_simp_tac (HOL_basic_ss addsimps [fresh_def]))) 1)]) - - val (_, thy10) = thy9 |> - Sign.add_path big_name |> - PureThy.add_thms [((Binding.name "strong_induct'", induct_aux), [])] ||>> - PureThy.add_thms [((Binding.name "strong_induct", induct), [case_names_induct])] ||>> - PureThy.add_thmss [((Binding.name "strong_inducts", projections induct), [case_names_induct])]; - - (**** recursion combinator ****) - - val _ = warning "defining recursion combinator ..."; - - val used = List.foldr OldTerm.add_typ_tfree_names [] recTs; - - val (rec_result_Ts', rec_fn_Ts') = DatatypeProp.make_primrec_Ts descr' sorts used; - - val rec_sort = if null dt_atomTs then HOLogic.typeS else - Sign.certify_sort thy10 pt_cp_sort; - - val rec_result_Ts = map (fn TFree (s, _) => TFree (s, rec_sort)) rec_result_Ts'; - val rec_fn_Ts = map (typ_subst_atomic (rec_result_Ts' ~~ rec_result_Ts)) rec_fn_Ts'; - - val rec_set_Ts = map (fn (T1, T2) => - rec_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts); - - val big_rec_name = big_name ^ "_rec_set"; - val rec_set_names' = - if length descr'' = 1 then [big_rec_name] else - map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int) - (1 upto (length descr'')); - val rec_set_names = map (Sign.full_bname thy10) rec_set_names'; - - val rec_fns = map (uncurry (mk_Free "f")) - (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts))); - val rec_sets' = map (fn c => list_comb (Free c, rec_fns)) - (rec_set_names' ~~ rec_set_Ts); - val rec_sets = map (fn c => list_comb (Const c, rec_fns)) - (rec_set_names ~~ rec_set_Ts); - - (* introduction rules for graph of recursion function *) - - val rec_preds = map (fn (a, T) => - Free (a, T --> HOLogic.boolT)) (pnames ~~ rec_result_Ts); - - fun mk_fresh3 rs [] = [] - | mk_fresh3 rs ((p as (ys, z)) :: yss) = List.concat (map (fn y as (_, T) => - List.mapPartial (fn ((_, (_, x)), r as (_, U)) => if z = x then NONE - else SOME (HOLogic.mk_Trueprop - (fresh_const T U $ Free y $ Free r))) rs) ys) @ - mk_fresh3 rs yss; - - (* FIXME: avoid collisions with other variable names? *) - val rec_ctxt = Free ("z", fsT'); - - fun make_rec_intr T p rec_set ((rec_intr_ts, rec_prems, rec_prems', - rec_eq_prems, l), ((cname, cargs), idxs)) = - let - val Ts = map (typ_of_dtyp descr'' sorts) cargs; - val frees = map (fn i => "x" ^ string_of_int i) (1 upto length Ts) ~~ Ts; - val frees' = partition_cargs idxs frees; - val binders = List.concat (map fst frees'); - val atomTs = distinct op = (maps (map snd o fst) frees'); - val recs = List.mapPartial - (fn ((_, DtRec i), p) => SOME (i, p) | _ => NONE) - (partition_cargs idxs cargs ~~ frees'); - val frees'' = map (fn i => "y" ^ string_of_int i) (1 upto length recs) ~~ - map (fn (i, _) => List.nth (rec_result_Ts, i)) recs; - val prems1 = map (fn ((i, (_, x)), y) => HOLogic.mk_Trueprop - (List.nth (rec_sets', i) $ Free x $ Free y)) (recs ~~ frees''); - val prems2 = - map (fn f => map (fn p as (_, T) => HOLogic.mk_Trueprop - (fresh_const T (fastype_of f) $ Free p $ f)) binders) rec_fns; - val prems3 = mk_fresh1 [] binders @ mk_fresh2 [] frees'; - val prems4 = map (fn ((i, _), y) => - HOLogic.mk_Trueprop (List.nth (rec_preds, i) $ Free y)) (recs ~~ frees''); - val prems5 = mk_fresh3 (recs ~~ frees'') frees'; - val prems6 = maps (fn aT => map (fn y as (_, T) => HOLogic.mk_Trueprop - (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $ - (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ Free y))) - frees'') atomTs; - val prems7 = map (fn x as (_, T) => HOLogic.mk_Trueprop - (fresh_const T fsT' $ Free x $ rec_ctxt)) binders; - val result = list_comb (List.nth (rec_fns, l), map Free (frees @ frees'')); - val result_freshs = map (fn p as (_, T) => - fresh_const T (fastype_of result) $ Free p $ result) binders; - val P = HOLogic.mk_Trueprop (p $ result) - in - (rec_intr_ts @ [Logic.list_implies (List.concat prems2 @ prems3 @ prems1, - HOLogic.mk_Trueprop (rec_set $ - list_comb (Const (cname, Ts ---> T), map Free frees) $ result))], - rec_prems @ [list_all_free (frees @ frees'', Logic.list_implies (prems4, P))], - rec_prems' @ map (fn fr => list_all_free (frees @ frees'', - Logic.list_implies (List.nth (prems2, l) @ prems3 @ prems5 @ prems7 @ prems6 @ [P], - HOLogic.mk_Trueprop fr))) result_freshs, - rec_eq_prems @ [List.concat prems2 @ prems3], - l + 1) - end; - - val (rec_intr_ts, rec_prems, rec_prems', rec_eq_prems, _) = - Library.foldl (fn (x, ((((d, d'), T), p), rec_set)) => - Library.foldl (make_rec_intr T p rec_set) (x, #3 (snd d) ~~ snd d')) - (([], [], [], [], 0), descr'' ~~ ndescr ~~ recTs ~~ rec_preds ~~ rec_sets'); - - val ({intrs = rec_intrs, elims = rec_elims, raw_induct = rec_induct, ...}, thy11) = - thy10 |> - Inductive.add_inductive_global (serial_string ()) - {quiet_mode = #quiet config, verbose = false, kind = Thm.internalK, - alt_name = Binding.name big_rec_name, coind = false, no_elim = false, no_ind = false, - skip_mono = true, fork_mono = false} - (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts)) - (map dest_Free rec_fns) - (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] ||> - PureThy.hide_fact true (Long_Name.append (Sign.full_bname thy10 big_rec_name) "induct"); - - (** equivariance **) - - val fresh_bij = PureThy.get_thms thy11 "fresh_bij"; - val perm_bij = PureThy.get_thms thy11 "perm_bij"; - - val (rec_equiv_thms, rec_equiv_thms') = ListPair.unzip (map (fn aT => - let - val permT = mk_permT aT; - val pi = Free ("pi", permT); - val rec_fns_pi = map (mk_perm [] pi o uncurry (mk_Free "f")) - (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts))); - val rec_sets_pi = map (fn c => list_comb (Const c, rec_fns_pi)) - (rec_set_names ~~ rec_set_Ts); - val ps = map (fn ((((T, U), R), R'), i) => - let - val x = Free ("x" ^ string_of_int i, T); - val y = Free ("y" ^ string_of_int i, U) - in - (R $ x $ y, R' $ mk_perm [] pi x $ mk_perm [] pi y) - end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ rec_sets_pi ~~ (1 upto length recTs)); - val ths = map (fn th => standard (th RS mp)) (split_conj_thm - (Goal.prove_global thy11 [] [] - (augment_sort thy1 pt_cp_sort - (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps)))) - (fn _ => rtac rec_induct 1 THEN REPEAT - (simp_tac (Simplifier.theory_context thy11 HOL_basic_ss - addsimps flat perm_simps' - addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN - (resolve_tac rec_intrs THEN_ALL_NEW - asm_simp_tac (HOL_ss addsimps (fresh_bij @ perm_bij))) 1)))) - val ths' = map (fn ((P, Q), th) => - Goal.prove_global thy11 [] [] - (augment_sort thy1 pt_cp_sort - (Logic.mk_implies (HOLogic.mk_Trueprop Q, HOLogic.mk_Trueprop P))) - (fn _ => dtac (Thm.instantiate ([], - [(cterm_of thy11 (Var (("pi", 0), permT)), - cterm_of thy11 (Const ("List.rev", permT --> permT) $ pi))]) th) 1 THEN - NominalPermeq.perm_simp_tac HOL_ss 1)) (ps ~~ ths) - in (ths, ths') end) dt_atomTs); - - (** finite support **) - - val rec_fin_supp_thms = map (fn aT => - let - val name = Long_Name.base_name (fst (dest_Type aT)); - val fs_name = PureThy.get_thm thy11 ("fs_" ^ name ^ "1"); - val aset = HOLogic.mk_setT aT; - val finite = Const ("Finite_Set.finite", aset --> HOLogic.boolT); - val fins = map (fn (f, T) => HOLogic.mk_Trueprop - (finite $ (Const ("Nominal.supp", T --> aset) $ f))) - (rec_fns ~~ rec_fn_Ts) - in - map (fn th => standard (th RS mp)) (split_conj_thm - (Goal.prove_global thy11 [] - (map (augment_sort thy11 fs_cp_sort) fins) - (augment_sort thy11 fs_cp_sort - (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj - (map (fn (((T, U), R), i) => - let - val x = Free ("x" ^ string_of_int i, T); - val y = Free ("y" ^ string_of_int i, U) - in - HOLogic.mk_imp (R $ x $ y, - finite $ (Const ("Nominal.supp", U --> aset) $ y)) - end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ - (1 upto length recTs)))))) - (fn {prems = fins, ...} => - (rtac rec_induct THEN_ALL_NEW cut_facts_tac fins) 1 THEN REPEAT - (NominalPermeq.finite_guess_tac (HOL_ss addsimps [fs_name]) 1)))) - end) dt_atomTs; - - (** freshness **) - - val finite_premss = map (fn aT => - map (fn (f, T) => HOLogic.mk_Trueprop - (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $ - (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ f))) - (rec_fns ~~ rec_fn_Ts)) dt_atomTs; - - val rec_fns' = map (augment_sort thy11 fs_cp_sort) rec_fns; - - val rec_fresh_thms = map (fn ((aT, eqvt_ths), finite_prems) => - let - val name = Long_Name.base_name (fst (dest_Type aT)); - val fs_name = PureThy.get_thm thy11 ("fs_" ^ name ^ "1"); - val a = Free ("a", aT); - val freshs = map (fn (f, fT) => HOLogic.mk_Trueprop - (fresh_const aT fT $ a $ f)) (rec_fns ~~ rec_fn_Ts) - in - map (fn (((T, U), R), eqvt_th) => - let - val x = Free ("x", augment_sort_typ thy11 fs_cp_sort T); - val y = Free ("y", U); - val y' = Free ("y'", U) - in - standard (Goal.prove (ProofContext.init thy11) [] - (map (augment_sort thy11 fs_cp_sort) - (finite_prems @ - [HOLogic.mk_Trueprop (R $ x $ y), - HOLogic.mk_Trueprop (HOLogic.mk_all ("y'", U, - HOLogic.mk_imp (R $ x $ y', HOLogic.mk_eq (y', y)))), - HOLogic.mk_Trueprop (fresh_const aT T $ a $ x)] @ - freshs)) - (HOLogic.mk_Trueprop (fresh_const aT U $ a $ y)) - (fn {prems, context} => - let - val (finite_prems, rec_prem :: unique_prem :: - fresh_prems) = chop (length finite_prems) prems; - val unique_prem' = unique_prem RS spec RS mp; - val unique = [unique_prem', unique_prem' RS sym] MRS trans; - val _ $ (_ $ (_ $ S $ _)) $ _ = prop_of supports_fresh; - val tuple = foldr1 HOLogic.mk_prod (x :: rec_fns') - in EVERY - [rtac (Drule.cterm_instantiate - [(cterm_of thy11 S, - cterm_of thy11 (Const ("Nominal.supp", - fastype_of tuple --> HOLogic.mk_setT aT) $ tuple))] - supports_fresh) 1, - simp_tac (HOL_basic_ss addsimps - [supports_def, symmetric fresh_def, fresh_prod]) 1, - REPEAT_DETERM (resolve_tac [allI, impI] 1), - REPEAT_DETERM (etac conjE 1), - rtac unique 1, - SUBPROOF (fn {prems = prems', params = [a, b], ...} => EVERY - [cut_facts_tac [rec_prem] 1, - rtac (Thm.instantiate ([], - [(cterm_of thy11 (Var (("pi", 0), mk_permT aT)), - cterm_of thy11 (perm_of_pair (term_of a, term_of b)))]) eqvt_th) 1, - asm_simp_tac (HOL_ss addsimps - (prems' @ perm_swap @ perm_fresh_fresh)) 1]) context 1, - rtac rec_prem 1, - simp_tac (HOL_ss addsimps (fs_name :: - supp_prod :: finite_Un :: finite_prems)) 1, - simp_tac (HOL_ss addsimps (symmetric fresh_def :: - fresh_prod :: fresh_prems)) 1] - end)) - end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ eqvt_ths) - end) (dt_atomTs ~~ rec_equiv_thms' ~~ finite_premss); - - (** uniqueness **) - - val fun_tuple = foldr1 HOLogic.mk_prod (rec_ctxt :: rec_fns); - val fun_tupleT = fastype_of fun_tuple; - val rec_unique_frees = - DatatypeProp.indexify_names (replicate (length recTs) "x") ~~ recTs; - val rec_unique_frees'' = map (fn (s, T) => (s ^ "'", T)) rec_unique_frees; - val rec_unique_frees' = - DatatypeProp.indexify_names (replicate (length recTs) "y") ~~ rec_result_Ts; - val rec_unique_concls = map (fn ((x, U), R) => - Const ("Ex1", (U --> HOLogic.boolT) --> HOLogic.boolT) $ - Abs ("y", U, R $ Free x $ Bound 0)) - (rec_unique_frees ~~ rec_result_Ts ~~ rec_sets); - - val induct_aux_rec = Drule.cterm_instantiate - (map (pairself (cterm_of thy11) o apsnd (augment_sort thy11 fs_cp_sort)) - (map (fn (aT, f) => (Logic.varify f, Abs ("z", HOLogic.unitT, - Const ("Nominal.supp", fun_tupleT --> HOLogic.mk_setT aT) $ fun_tuple))) - fresh_fs @ - map (fn (((P, T), (x, U)), Q) => - (Var ((P, 0), Logic.varifyT (fsT' --> T --> HOLogic.boolT)), - Abs ("z", HOLogic.unitT, absfree (x, U, Q)))) - (pnames ~~ recTs ~~ rec_unique_frees ~~ rec_unique_concls) @ - map (fn (s, T) => (Var ((s, 0), Logic.varifyT T), Free (s, T))) - rec_unique_frees)) induct_aux; - - fun obtain_fresh_name vs ths rec_fin_supp T (freshs1, freshs2, ctxt) = - let - val p = foldr1 HOLogic.mk_prod (vs @ freshs1); - val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop - (HOLogic.exists_const T $ Abs ("x", T, - fresh_const T (fastype_of p) $ Bound 0 $ p))) - (fn _ => EVERY - [cut_facts_tac ths 1, - REPEAT_DETERM (dresolve_tac (the (AList.lookup op = rec_fin_supp T)) 1), - resolve_tac exists_fresh' 1, - asm_simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms)) 1]); - val (([cx], ths), ctxt') = Obtain.result - (fn _ => EVERY - [etac exE 1, - full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1, - REPEAT (etac conjE 1)]) - [ex] ctxt - in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end; - - val finite_ctxt_prems = map (fn aT => - HOLogic.mk_Trueprop - (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $ - (Const ("Nominal.supp", fsT' --> HOLogic.mk_setT aT) $ rec_ctxt))) dt_atomTs; - - val rec_unique_thms = split_conj_thm (Goal.prove - (ProofContext.init thy11) (map fst rec_unique_frees) - (map (augment_sort thy11 fs_cp_sort) - (List.concat finite_premss @ finite_ctxt_prems @ rec_prems @ rec_prems')) - (augment_sort thy11 fs_cp_sort - (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj rec_unique_concls))) - (fn {prems, context} => - let - val k = length rec_fns; - val (finite_thss, ths1) = fold_map (fn T => fn xs => - apfst (pair T) (chop k xs)) dt_atomTs prems; - val (finite_ctxt_ths, ths2) = chop (length dt_atomTs) ths1; - val (P_ind_ths, fcbs) = chop k ths2; - val P_ths = map (fn th => th RS mp) (split_conj_thm - (Goal.prove context - (map fst (rec_unique_frees'' @ rec_unique_frees')) [] - (augment_sort thy11 fs_cp_sort - (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj - (map (fn (((x, y), S), P) => HOLogic.mk_imp - (S $ Free x $ Free y, P $ (Free y))) - (rec_unique_frees'' ~~ rec_unique_frees' ~~ - rec_sets ~~ rec_preds))))) - (fn _ => - rtac rec_induct 1 THEN - REPEAT ((resolve_tac P_ind_ths THEN_ALL_NEW assume_tac) 1)))); - val rec_fin_supp_thms' = map - (fn (ths, (T, fin_ths)) => (T, map (curry op MRS fin_ths) ths)) - (rec_fin_supp_thms ~~ finite_thss); - in EVERY - ([rtac induct_aux_rec 1] @ - maps (fn ((_, finite_ths), finite_th) => - [cut_facts_tac (finite_th :: finite_ths) 1, - asm_simp_tac (HOL_ss addsimps [supp_prod, finite_Un]) 1]) - (finite_thss ~~ finite_ctxt_ths) @ - maps (fn ((_, idxss), elim) => maps (fn idxs => - [full_simp_tac (HOL_ss addsimps [symmetric fresh_def, supp_prod, Un_iff]) 1, - REPEAT_DETERM (eresolve_tac [conjE, ex1E] 1), - rtac ex1I 1, - (resolve_tac rec_intrs THEN_ALL_NEW atac) 1, - rotate_tac ~1 1, - ((DETERM o etac elim) THEN_ALL_NEW full_simp_tac - (HOL_ss addsimps List.concat distinct_thms)) 1] @ - (if null idxs then [] else [hyp_subst_tac 1, - SUBPROOF (fn {asms, concl, prems = prems', params, context = context', ...} => - let - val SOME prem = find_first (can (HOLogic.dest_eq o - HOLogic.dest_Trueprop o prop_of)) prems'; - val _ $ (_ $ lhs $ rhs) = prop_of prem; - val _ $ (_ $ lhs' $ rhs') = term_of concl; - val rT = fastype_of lhs'; - val (c, cargsl) = strip_comb lhs; - val cargsl' = partition_cargs idxs cargsl; - val boundsl = List.concat (map fst cargsl'); - val (_, cargsr) = strip_comb rhs; - val cargsr' = partition_cargs idxs cargsr; - val boundsr = List.concat (map fst cargsr'); - val (params1, _ :: params2) = - chop (length params div 2) (map term_of params); - val params' = params1 @ params2; - val rec_prems = filter (fn th => case prop_of th of - _ $ p => (case head_of p of - Const (s, _) => s mem rec_set_names - | _ => false) - | _ => false) prems'; - val fresh_prems = filter (fn th => case prop_of th of - _ $ (Const ("Nominal.fresh", _) $ _ $ _) => true - | _ $ (Const ("Not", _) $ _) => true - | _ => false) prems'; - val Ts = map fastype_of boundsl; - - val _ = warning "step 1: obtaining fresh names"; - val (freshs1, freshs2, context'') = fold - (obtain_fresh_name (rec_ctxt :: rec_fns' @ params') - (List.concat (map snd finite_thss) @ - finite_ctxt_ths @ rec_prems) - rec_fin_supp_thms') - Ts ([], [], context'); - val pi1 = map perm_of_pair (boundsl ~~ freshs1); - val rpi1 = rev pi1; - val pi2 = map perm_of_pair (boundsr ~~ freshs1); - val rpi2 = rev pi2; - - val fresh_prems' = mk_not_sym fresh_prems; - val freshs2' = mk_not_sym freshs2; - - (** as, bs, cs # K as ts, K bs us **) - val _ = warning "step 2: as, bs, cs # K as ts, K bs us"; - val prove_fresh_ss = HOL_ss addsimps - (finite_Diff :: List.concat fresh_thms @ - fs_atoms @ abs_fresh @ abs_supp @ fresh_atm); - (* FIXME: avoid asm_full_simp_tac ? *) - fun prove_fresh ths y x = Goal.prove context'' [] [] - (HOLogic.mk_Trueprop (fresh_const - (fastype_of x) (fastype_of y) $ x $ y)) - (fn _ => cut_facts_tac ths 1 THEN asm_full_simp_tac prove_fresh_ss 1); - val constr_fresh_thms = - map (prove_fresh fresh_prems lhs) boundsl @ - map (prove_fresh fresh_prems rhs) boundsr @ - map (prove_fresh freshs2 lhs) freshs1 @ - map (prove_fresh freshs2 rhs) freshs1; - - (** pi1 o (K as ts) = pi2 o (K bs us) **) - val _ = warning "step 3: pi1 o (K as ts) = pi2 o (K bs us)"; - val pi1_pi2_eq = Goal.prove context'' [] [] - (HOLogic.mk_Trueprop (HOLogic.mk_eq - (fold_rev (mk_perm []) pi1 lhs, fold_rev (mk_perm []) pi2 rhs))) - (fn _ => EVERY - [cut_facts_tac constr_fresh_thms 1, - asm_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh) 1, - rtac prem 1]); - - (** pi1 o ts = pi2 o us **) - val _ = warning "step 4: pi1 o ts = pi2 o us"; - val pi1_pi2_eqs = map (fn (t, u) => - Goal.prove context'' [] [] - (HOLogic.mk_Trueprop (HOLogic.mk_eq - (fold_rev (mk_perm []) pi1 t, fold_rev (mk_perm []) pi2 u))) - (fn _ => EVERY - [cut_facts_tac [pi1_pi2_eq] 1, - asm_full_simp_tac (HOL_ss addsimps - (calc_atm @ List.concat perm_simps' @ - fresh_prems' @ freshs2' @ abs_perm @ - alpha @ List.concat inject_thms)) 1])) - (map snd cargsl' ~~ map snd cargsr'); - - (** pi1^-1 o pi2 o us = ts **) - val _ = warning "step 5: pi1^-1 o pi2 o us = ts"; - val rpi1_pi2_eqs = map (fn ((t, u), eq) => - Goal.prove context'' [] [] - (HOLogic.mk_Trueprop (HOLogic.mk_eq - (fold_rev (mk_perm []) (rpi1 @ pi2) u, t))) - (fn _ => simp_tac (HOL_ss addsimps - ((eq RS sym) :: perm_swap)) 1)) - (map snd cargsl' ~~ map snd cargsr' ~~ pi1_pi2_eqs); - - val (rec_prems1, rec_prems2) = - chop (length rec_prems div 2) rec_prems; - - (** (ts, pi1^-1 o pi2 o vs) in rec_set **) - val _ = warning "step 6: (ts, pi1^-1 o pi2 o vs) in rec_set"; - val rec_prems' = map (fn th => - let - val _ $ (S $ x $ y) = prop_of th; - val Const (s, _) = head_of S; - val k = find_index (equal s) rec_set_names; - val pi = rpi1 @ pi2; - fun mk_pi z = fold_rev (mk_perm []) pi z; - fun eqvt_tac p = - let - val U as Type (_, [Type (_, [T, _])]) = fastype_of p; - val l = find_index (equal T) dt_atomTs; - val th = List.nth (List.nth (rec_equiv_thms', l), k); - val th' = Thm.instantiate ([], - [(cterm_of thy11 (Var (("pi", 0), U)), - cterm_of thy11 p)]) th; - in rtac th' 1 end; - val th' = Goal.prove context'' [] [] - (HOLogic.mk_Trueprop (S $ mk_pi x $ mk_pi y)) - (fn _ => EVERY - (map eqvt_tac pi @ - [simp_tac (HOL_ss addsimps (fresh_prems' @ freshs2' @ - perm_swap @ perm_fresh_fresh)) 1, - rtac th 1])) - in - Simplifier.simplify - (HOL_basic_ss addsimps rpi1_pi2_eqs) th' - end) rec_prems2; - - val ihs = filter (fn th => case prop_of th of - _ $ (Const ("All", _) $ _) => true | _ => false) prems'; - - (** pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs **) - val _ = warning "step 7: pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs"; - val rec_eqns = map (fn (th, ih) => - let - val th' = th RS (ih RS spec RS mp) RS sym; - val _ $ (_ $ lhs $ rhs) = prop_of th'; - fun strip_perm (_ $ _ $ t) = strip_perm t - | strip_perm t = t; - in - Goal.prove context'' [] [] - (HOLogic.mk_Trueprop (HOLogic.mk_eq - (fold_rev (mk_perm []) pi1 lhs, - fold_rev (mk_perm []) pi2 (strip_perm rhs)))) - (fn _ => simp_tac (HOL_basic_ss addsimps - (th' :: perm_swap)) 1) - end) (rec_prems' ~~ ihs); - - (** as # rs **) - val _ = warning "step 8: as # rs"; - val rec_freshs = List.concat - (map (fn (rec_prem, ih) => - let - val _ $ (S $ x $ (y as Free (_, T))) = - prop_of rec_prem; - val k = find_index (equal S) rec_sets; - val atoms = List.concat (List.mapPartial (fn (bs, z) => - if z = x then NONE else SOME bs) cargsl') - in - map (fn a as Free (_, aT) => - let val l = find_index (equal aT) dt_atomTs; - in - Goal.prove context'' [] [] - (HOLogic.mk_Trueprop (fresh_const aT T $ a $ y)) - (fn _ => EVERY - (rtac (List.nth (List.nth (rec_fresh_thms, l), k)) 1 :: - map (fn th => rtac th 1) - (snd (List.nth (finite_thss, l))) @ - [rtac rec_prem 1, rtac ih 1, - REPEAT_DETERM (resolve_tac fresh_prems 1)])) - end) atoms - end) (rec_prems1 ~~ ihs)); - - (** as # fK as ts rs , bs # fK bs us vs **) - val _ = warning "step 9: as # fK as ts rs , bs # fK bs us vs"; - fun prove_fresh_result (a as Free (_, aT)) = - Goal.prove context'' [] [] - (HOLogic.mk_Trueprop (fresh_const aT rT $ a $ rhs')) - (fn _ => EVERY - [resolve_tac fcbs 1, - REPEAT_DETERM (resolve_tac - (fresh_prems @ rec_freshs) 1), - REPEAT_DETERM (resolve_tac (maps snd rec_fin_supp_thms') 1 - THEN resolve_tac rec_prems 1), - resolve_tac P_ind_ths 1, - REPEAT_DETERM (resolve_tac (P_ths @ rec_prems) 1)]); - - val fresh_results'' = map prove_fresh_result boundsl; - - fun prove_fresh_result'' ((a as Free (_, aT), b), th) = - let val th' = Goal.prove context'' [] [] - (HOLogic.mk_Trueprop (fresh_const aT rT $ - fold_rev (mk_perm []) (rpi2 @ pi1) a $ - fold_rev (mk_perm []) (rpi2 @ pi1) rhs')) - (fn _ => simp_tac (HOL_ss addsimps fresh_bij) 1 THEN - rtac th 1) - in - Goal.prove context'' [] [] - (HOLogic.mk_Trueprop (fresh_const aT rT $ b $ lhs')) - (fn _ => EVERY - [cut_facts_tac [th'] 1, - full_simp_tac (Simplifier.theory_context thy11 HOL_ss - addsimps rec_eqns @ pi1_pi2_eqs @ perm_swap - addsimprocs [NominalPermeq.perm_simproc_app]) 1, - full_simp_tac (HOL_ss addsimps (calc_atm @ - fresh_prems' @ freshs2' @ perm_fresh_fresh)) 1]) - end; - - val fresh_results = fresh_results'' @ map prove_fresh_result'' - (boundsl ~~ boundsr ~~ fresh_results''); - - (** cs # fK as ts rs , cs # fK bs us vs **) - val _ = warning "step 10: cs # fK as ts rs , cs # fK bs us vs"; - fun prove_fresh_result' recs t (a as Free (_, aT)) = - Goal.prove context'' [] [] - (HOLogic.mk_Trueprop (fresh_const aT rT $ a $ t)) - (fn _ => EVERY - [cut_facts_tac recs 1, - REPEAT_DETERM (dresolve_tac - (the (AList.lookup op = rec_fin_supp_thms' aT)) 1), - NominalPermeq.fresh_guess_tac - (HOL_ss addsimps (freshs2 @ - fs_atoms @ fresh_atm @ - List.concat (map snd finite_thss))) 1]); - - val fresh_results' = - map (prove_fresh_result' rec_prems1 rhs') freshs1 @ - map (prove_fresh_result' rec_prems2 lhs') freshs1; - - (** pi1 o (fK as ts rs) = pi2 o (fK bs us vs) **) - val _ = warning "step 11: pi1 o (fK as ts rs) = pi2 o (fK bs us vs)"; - val pi1_pi2_result = Goal.prove context'' [] [] - (HOLogic.mk_Trueprop (HOLogic.mk_eq - (fold_rev (mk_perm []) pi1 rhs', fold_rev (mk_perm []) pi2 lhs'))) - (fn _ => simp_tac (Simplifier.context context'' HOL_ss - addsimps pi1_pi2_eqs @ rec_eqns - addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN - TRY (simp_tac (HOL_ss addsimps - (fresh_prems' @ freshs2' @ calc_atm @ perm_fresh_fresh)) 1)); - - val _ = warning "final result"; - val final = Goal.prove context'' [] [] (term_of concl) - (fn _ => cut_facts_tac [pi1_pi2_result RS sym] 1 THEN - full_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh @ - fresh_results @ fresh_results') 1); - val final' = ProofContext.export context'' context' [final]; - val _ = warning "finished!" - in - resolve_tac final' 1 - end) context 1])) idxss) (ndescr ~~ rec_elims)) - end)); - - val rec_total_thms = map (fn r => r RS theI') rec_unique_thms; - - (* define primrec combinators *) - - val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec"; - val reccomb_names = map (Sign.full_bname thy11) - (if length descr'' = 1 then [big_reccomb_name] else - (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int) - (1 upto (length descr'')))); - val reccombs = map (fn ((name, T), T') => list_comb - (Const (name, rec_fn_Ts @ [T] ---> T'), rec_fns)) - (reccomb_names ~~ recTs ~~ rec_result_Ts); - - val (reccomb_defs, thy12) = - thy11 - |> Sign.add_consts_i (map (fn ((name, T), T') => - (Binding.name (Long_Name.base_name name), rec_fn_Ts @ [T] ---> T', NoSyn)) - (reccomb_names ~~ recTs ~~ rec_result_Ts)) - |> (PureThy.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') => - (Binding.name (Long_Name.base_name name ^ "_def"), Logic.mk_equals (comb, absfree ("x", T, - Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T', - set $ Free ("x", T) $ Free ("y", T')))))) - (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts)); - - (* prove characteristic equations for primrec combinators *) - - val rec_thms = map (fn (prems, concl) => - let - val _ $ (_ $ (_ $ x) $ _) = concl; - val (_, cargs) = strip_comb x; - val ps = map (fn (x as Free (_, T), i) => - (Free ("x" ^ string_of_int i, T), x)) (cargs ~~ (1 upto length cargs)); - val concl' = subst_atomic_types (rec_result_Ts' ~~ rec_result_Ts) concl; - val prems' = List.concat finite_premss @ finite_ctxt_prems @ - rec_prems @ rec_prems' @ map (subst_atomic ps) prems; - fun solve rules prems = resolve_tac rules THEN_ALL_NEW - (resolve_tac prems THEN_ALL_NEW atac) - in - Goal.prove_global thy12 [] - (map (augment_sort thy12 fs_cp_sort) prems') - (augment_sort thy12 fs_cp_sort concl') - (fn {prems, ...} => EVERY - [rewrite_goals_tac reccomb_defs, - rtac the1_equality 1, - solve rec_unique_thms prems 1, - resolve_tac rec_intrs 1, - REPEAT (solve (prems @ rec_total_thms) prems 1)]) - end) (rec_eq_prems ~~ - DatatypeProp.make_primrecs new_type_names descr' sorts thy12); - - val dt_infos = map (make_dt_info pdescr sorts induct reccomb_names rec_thms) - ((0 upto length descr1 - 1) ~~ descr1 ~~ distinct_thms ~~ inject_thms); - - (* FIXME: theorems are stored in database for testing only *) - val (_, thy13) = thy12 |> - PureThy.add_thmss - [((Binding.name "rec_equiv", List.concat rec_equiv_thms), []), - ((Binding.name "rec_equiv'", List.concat rec_equiv_thms'), []), - ((Binding.name "rec_fin_supp", List.concat rec_fin_supp_thms), []), - ((Binding.name "rec_fresh", List.concat rec_fresh_thms), []), - ((Binding.name "rec_unique", map standard rec_unique_thms), []), - ((Binding.name "recs", rec_thms), [])] ||> - Sign.parent_path ||> - map_nominal_datatypes (fold Symtab.update dt_infos); - - in - thy13 - end; - -val add_nominal_datatype = gen_add_nominal_datatype Datatype.read_typ; - - -(* FIXME: The following stuff should be exported by Datatype *) - -local structure P = OuterParse and K = OuterKeyword in - -val datatype_decl = - Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix -- - (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix)); - -fun mk_datatype args = - let - val names = map (fn ((((NONE, _), t), _), _) => t | ((((SOME t, _), _), _), _) => t) args; - val specs = map (fn ((((_, vs), t), mx), cons) => - (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args; - in add_nominal_datatype Datatype.default_config names specs end; - -val _ = - OuterSyntax.command "nominal_datatype" "define inductive datatypes" K.thy_decl - (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype)); - -end; - -end diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Nominal/nominal_datatype.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nominal/nominal_datatype.ML Sat Jul 04 15:19:47 2009 +0200 @@ -0,0 +1,2094 @@ +(* Title: HOL/Nominal/nominal_datatype.ML + Author: Stefan Berghofer and Christian Urban, TU Muenchen + +Nominal datatype package for Isabelle/HOL. +*) + +signature NOMINAL_DATATYPE = +sig + val add_nominal_datatype : Datatype.config -> string list -> + (string list * bstring * mixfix * + (bstring * string list * mixfix) list) list -> theory -> theory + type descr + type nominal_datatype_info + val get_nominal_datatypes : theory -> nominal_datatype_info Symtab.table + val get_nominal_datatype : theory -> string -> nominal_datatype_info option + val mk_perm: typ list -> term -> term -> term + val perm_of_pair: term * term -> term + val mk_not_sym: thm list -> thm list + val perm_simproc: simproc + val fresh_const: typ -> typ -> term + val fresh_star_const: typ -> typ -> term +end + +structure NominalDatatype : NOMINAL_DATATYPE = +struct + +val finite_emptyI = thm "finite.emptyI"; +val finite_Diff = thm "finite_Diff"; +val finite_Un = thm "finite_Un"; +val Un_iff = thm "Un_iff"; +val In0_eq = thm "In0_eq"; +val In1_eq = thm "In1_eq"; +val In0_not_In1 = thm "In0_not_In1"; +val In1_not_In0 = thm "In1_not_In0"; +val Un_assoc = thm "Un_assoc"; +val Collect_disj_eq = thm "Collect_disj_eq"; +val empty_def = thm "empty_def"; +val empty_iff = thm "empty_iff"; + +open DatatypeAux; +open NominalAtoms; + +(** FIXME: Datatype should export this function **) + +local + +fun dt_recs (DtTFree _) = [] + | dt_recs (DtType (_, dts)) = List.concat (map dt_recs dts) + | dt_recs (DtRec i) = [i]; + +fun dt_cases (descr: descr) (_, args, constrs) = + let + fun the_bname i = Long_Name.base_name (#1 (valOf (AList.lookup (op =) descr i))); + val bnames = map the_bname (distinct op = (List.concat (map dt_recs args))); + in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end; + + +fun induct_cases descr = + DatatypeProp.indexify_names (List.concat (map (dt_cases descr) (map #2 descr))); + +fun exhaust_cases descr i = dt_cases descr (valOf (AList.lookup (op =) descr i)); + +in + +fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr); + +fun mk_case_names_exhausts descr new = + map (RuleCases.case_names o exhaust_cases descr o #1) + (List.filter (fn ((_, (name, _, _))) => name mem_string new) descr); + +end; + +(* theory data *) + +type descr = (int * (string * dtyp list * (string * (dtyp list * dtyp) list) list)) list; + +type nominal_datatype_info = + {index : int, + descr : descr, + sorts : (string * sort) list, + rec_names : string list, + rec_rewrites : thm list, + induction : thm, + distinct : thm list, + inject : thm list}; + +structure NominalDatatypesData = TheoryDataFun +( + type T = nominal_datatype_info Symtab.table; + val empty = Symtab.empty; + val copy = I; + val extend = I; + fun merge _ tabs : T = Symtab.merge (K true) tabs; +); + +val get_nominal_datatypes = NominalDatatypesData.get; +val put_nominal_datatypes = NominalDatatypesData.put; +val map_nominal_datatypes = NominalDatatypesData.map; +val get_nominal_datatype = Symtab.lookup o get_nominal_datatypes; + + +(**** make datatype info ****) + +fun make_dt_info descr sorts induct reccomb_names rec_thms + (((i, (_, (tname, _, _))), distinct), inject) = + (tname, + {index = i, + descr = descr, + sorts = sorts, + rec_names = reccomb_names, + rec_rewrites = rec_thms, + induction = induct, + distinct = distinct, + inject = inject}); + +(*******************************) + +val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma); + + +(** simplification procedure for sorting permutations **) + +val dj_cp = thm "dj_cp"; + +fun dest_permT (Type ("fun", [Type ("List.list", [Type ("*", [T, _])]), + Type ("fun", [_, U])])) = (T, U); + +fun permTs_of (Const ("Nominal.perm", T) $ t $ u) = fst (dest_permT T) :: permTs_of u + | permTs_of _ = []; + +fun perm_simproc' thy ss (Const ("Nominal.perm", T) $ t $ (u as Const ("Nominal.perm", U) $ r $ s)) = + let + val (aT as Type (a, []), S) = dest_permT T; + val (bT as Type (b, []), _) = dest_permT U + in if aT mem permTs_of u andalso aT <> bT then + let + val cp = cp_inst_of thy a b; + val dj = dj_thm_of thy b a; + val dj_cp' = [cp, dj] MRS dj_cp; + val cert = SOME o cterm_of thy + in + SOME (mk_meta_eq (Drule.instantiate' [SOME (ctyp_of thy S)] + [cert t, cert r, cert s] dj_cp')) + end + else NONE + end + | perm_simproc' thy ss _ = NONE; + +val perm_simproc = + Simplifier.simproc (the_context ()) "perm_simp" ["pi1 \ (pi2 \ x)"] perm_simproc'; + +val meta_spec = thm "meta_spec"; + +fun projections rule = + ProjectRule.projections (ProofContext.init (Thm.theory_of_thm rule)) rule + |> map (standard #> RuleCases.save rule); + +val supp_prod = thm "supp_prod"; +val fresh_prod = thm "fresh_prod"; +val supports_fresh = thm "supports_fresh"; +val supports_def = thm "Nominal.supports_def"; +val fresh_def = thm "fresh_def"; +val supp_def = thm "supp_def"; +val rev_simps = thms "rev.simps"; +val app_simps = thms "append.simps"; +val at_fin_set_supp = thm "at_fin_set_supp"; +val at_fin_set_fresh = thm "at_fin_set_fresh"; +val abs_fun_eq1 = thm "abs_fun_eq1"; + +val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq]; + +fun mk_perm Ts t u = + let + val T = fastype_of1 (Ts, t); + val U = fastype_of1 (Ts, u) + in Const ("Nominal.perm", T --> U --> U) $ t $ u end; + +fun perm_of_pair (x, y) = + let + val T = fastype_of x; + val pT = mk_permT T + in Const ("List.list.Cons", HOLogic.mk_prodT (T, T) --> pT --> pT) $ + HOLogic.mk_prod (x, y) $ Const ("List.list.Nil", pT) + end; + +fun mk_not_sym ths = maps (fn th => case prop_of th of + _ $ (Const ("Not", _) $ (Const ("op =", _) $ _ $ _)) => [th, th RS not_sym] + | _ => [th]) ths; + +fun fresh_const T U = Const ("Nominal.fresh", T --> U --> HOLogic.boolT); +fun fresh_star_const T U = + Const ("Nominal.fresh_star", HOLogic.mk_setT T --> U --> HOLogic.boolT); + +fun gen_add_nominal_datatype prep_typ config new_type_names dts thy = + let + (* this theory is used just for parsing *) + + val tmp_thy = thy |> + Theory.copy |> + Sign.add_types (map (fn (tvs, tname, mx, _) => + (Binding.name tname, length tvs, mx)) dts); + + val atoms = atoms_of thy; + + fun prep_constr ((constrs, sorts), (cname, cargs, mx)) = + let val (cargs', sorts') = Library.foldl (prep_typ tmp_thy) (([], sorts), cargs) + in (constrs @ [(cname, cargs', mx)], sorts') end + + fun prep_dt_spec ((dts, sorts), (tvs, tname, mx, constrs)) = + let val (constrs', sorts') = Library.foldl prep_constr (([], sorts), constrs) + in (dts @ [(tvs, tname, mx, constrs')], sorts') end + + val (dts', sorts) = Library.foldl prep_dt_spec (([], []), dts); + val tyvars = map (map (fn s => + (s, the (AList.lookup (op =) sorts s))) o #1) dts'; + + fun inter_sort thy S S' = Type.inter_sort (Sign.tsig_of thy) (S, S'); + fun augment_sort_typ thy S = + let val S = Sign.certify_sort thy S + in map_type_tfree (fn (s, S') => TFree (s, + if member (op = o apsnd fst) sorts s then inter_sort thy S S' else S')) + end; + fun augment_sort thy S = map_types (augment_sort_typ thy S); + + val types_syntax = map (fn (tvs, tname, mx, constrs) => (tname, mx)) dts'; + val constr_syntax = map (fn (tvs, tname, mx, constrs) => + map (fn (cname, cargs, mx) => (cname, mx)) constrs) dts'; + + val ps = map (fn (_, n, _, _) => + (Sign.full_bname tmp_thy n, Sign.full_bname tmp_thy (n ^ "_Rep"))) dts; + val rps = map Library.swap ps; + + fun replace_types (Type ("Nominal.ABS", [T, U])) = + Type ("fun", [T, Type ("Nominal.noption", [replace_types U])]) + | replace_types (Type (s, Ts)) = + Type (getOpt (AList.lookup op = ps s, s), map replace_types Ts) + | replace_types T = T; + + val dts'' = map (fn (tvs, tname, mx, constrs) => (tvs, Binding.name (tname ^ "_Rep"), NoSyn, + map (fn (cname, cargs, mx) => (Binding.name (cname ^ "_Rep"), + map replace_types cargs, NoSyn)) constrs)) dts'; + + val new_type_names' = map (fn n => n ^ "_Rep") new_type_names; + + val (full_new_type_names',thy1) = + Datatype.add_datatype config new_type_names' dts'' thy; + + val {descr, induction, ...} = + Datatype.the_info thy1 (hd full_new_type_names'); + fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); + + val big_name = space_implode "_" new_type_names; + + + (**** define permutation functions ****) + + val permT = mk_permT (TFree ("'x", HOLogic.typeS)); + val pi = Free ("pi", permT); + val perm_types = map (fn (i, _) => + let val T = nth_dtyp i + in permT --> T --> T end) descr; + val perm_names' = DatatypeProp.indexify_names (map (fn (i, _) => + "perm_" ^ name_of_typ (nth_dtyp i)) descr); + val perm_names = replicate (length new_type_names) "Nominal.perm" @ + map (Sign.full_bname thy1) (List.drop (perm_names', length new_type_names)); + val perm_names_types = perm_names ~~ perm_types; + val perm_names_types' = perm_names' ~~ perm_types; + + val perm_eqs = maps (fn (i, (_, _, constrs)) => + let val T = nth_dtyp i + in map (fn (cname, dts) => + let + val Ts = map (typ_of_dtyp descr sorts) dts; + val names = Name.variant_list ["pi"] (DatatypeProp.make_tnames Ts); + val args = map Free (names ~~ Ts); + val c = Const (cname, Ts ---> T); + fun perm_arg (dt, x) = + let val T = type_of x + in if is_rec_type dt then + let val (Us, _) = strip_type T + in list_abs (map (pair "x") Us, + Free (nth perm_names_types' (body_index dt)) $ pi $ + list_comb (x, map (fn (i, U) => + Const ("Nominal.perm", permT --> U --> U) $ + (Const ("List.rev", permT --> permT) $ pi) $ + Bound i) ((length Us - 1 downto 0) ~~ Us))) + end + else Const ("Nominal.perm", permT --> T --> T) $ pi $ x + end; + in + (Attrib.empty_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq + (Free (nth perm_names_types' i) $ + Free ("pi", mk_permT (TFree ("'x", HOLogic.typeS))) $ + list_comb (c, args), + list_comb (c, map perm_arg (dts ~~ args))))) + end) constrs + end) descr; + + val (perm_simps, thy2) = + Primrec.add_primrec_overloaded + (map (fn (s, sT) => (s, sT, false)) + (List.take (perm_names' ~~ perm_names_types, length new_type_names))) + (map (fn s => (Binding.name s, NONE, NoSyn)) perm_names') perm_eqs thy1; + + (**** prove that permutation functions introduced by unfolding are ****) + (**** equivalent to already existing permutation functions ****) + + val _ = warning ("length descr: " ^ string_of_int (length descr)); + val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names)); + + val perm_indnames = DatatypeProp.make_tnames (map body_type perm_types); + val perm_fun_def = PureThy.get_thm thy2 "perm_fun_def"; + + val unfolded_perm_eq_thms = + if length descr = length new_type_names then [] + else map standard (List.drop (split_conj_thm + (Goal.prove_global thy2 [] [] + (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj + (map (fn (c as (s, T), x) => + let val [T1, T2] = binder_types T + in HOLogic.mk_eq (Const c $ pi $ Free (x, T2), + Const ("Nominal.perm", T) $ pi $ Free (x, T2)) + end) + (perm_names_types ~~ perm_indnames)))) + (fn _ => EVERY [indtac induction perm_indnames 1, + ALLGOALS (asm_full_simp_tac + (simpset_of thy2 addsimps [perm_fun_def]))])), + length new_type_names)); + + (**** prove [] \ t = t ****) + + val _ = warning "perm_empty_thms"; + + val perm_empty_thms = List.concat (map (fn a => + let val permT = mk_permT (Type (a, [])) + in map standard (List.take (split_conj_thm + (Goal.prove_global thy2 [] [] + (augment_sort thy2 [pt_class_of thy2 a] + (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj + (map (fn ((s, T), x) => HOLogic.mk_eq + (Const (s, permT --> T --> T) $ + Const ("List.list.Nil", permT) $ Free (x, T), + Free (x, T))) + (perm_names ~~ + map body_type perm_types ~~ perm_indnames))))) + (fn _ => EVERY [indtac induction perm_indnames 1, + ALLGOALS (asm_full_simp_tac (simpset_of thy2))])), + length new_type_names)) + end) + atoms); + + (**** prove (pi1 @ pi2) \ t = pi1 \ (pi2 \ t) ****) + + val _ = warning "perm_append_thms"; + + (*FIXME: these should be looked up statically*) + val at_pt_inst = PureThy.get_thm thy2 "at_pt_inst"; + val pt2 = PureThy.get_thm thy2 "pt2"; + + val perm_append_thms = List.concat (map (fn a => + let + val permT = mk_permT (Type (a, [])); + val pi1 = Free ("pi1", permT); + val pi2 = Free ("pi2", permT); + val pt_inst = pt_inst_of thy2 a; + val pt2' = pt_inst RS pt2; + val pt2_ax = PureThy.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "2") a); + in List.take (map standard (split_conj_thm + (Goal.prove_global thy2 [] [] + (augment_sort thy2 [pt_class_of thy2 a] + (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj + (map (fn ((s, T), x) => + let val perm = Const (s, permT --> T --> T) + in HOLogic.mk_eq + (perm $ (Const ("List.append", permT --> permT --> permT) $ + pi1 $ pi2) $ Free (x, T), + perm $ pi1 $ (perm $ pi2 $ Free (x, T))) + end) + (perm_names ~~ + map body_type perm_types ~~ perm_indnames))))) + (fn _ => EVERY [indtac induction perm_indnames 1, + ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt2', pt2_ax]))]))), + length new_type_names) + end) atoms); + + (**** prove pi1 ~ pi2 ==> pi1 \ t = pi2 \ t ****) + + val _ = warning "perm_eq_thms"; + + val pt3 = PureThy.get_thm thy2 "pt3"; + val pt3_rev = PureThy.get_thm thy2 "pt3_rev"; + + val perm_eq_thms = List.concat (map (fn a => + let + val permT = mk_permT (Type (a, [])); + val pi1 = Free ("pi1", permT); + val pi2 = Free ("pi2", permT); + val at_inst = at_inst_of thy2 a; + val pt_inst = pt_inst_of thy2 a; + val pt3' = pt_inst RS pt3; + val pt3_rev' = at_inst RS (pt_inst RS pt3_rev); + val pt3_ax = PureThy.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "3") a); + in List.take (map standard (split_conj_thm + (Goal.prove_global thy2 [] [] + (augment_sort thy2 [pt_class_of thy2 a] (Logic.mk_implies + (HOLogic.mk_Trueprop (Const ("Nominal.prm_eq", + permT --> permT --> HOLogic.boolT) $ pi1 $ pi2), + HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj + (map (fn ((s, T), x) => + let val perm = Const (s, permT --> T --> T) + in HOLogic.mk_eq + (perm $ pi1 $ Free (x, T), + perm $ pi2 $ Free (x, T)) + end) + (perm_names ~~ + map body_type perm_types ~~ perm_indnames)))))) + (fn _ => EVERY [indtac induction perm_indnames 1, + ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt3', pt3_rev', pt3_ax]))]))), + length new_type_names) + end) atoms); + + (**** prove pi1 \ (pi2 \ t) = (pi1 \ pi2) \ (pi1 \ t) ****) + + val cp1 = PureThy.get_thm thy2 "cp1"; + val dj_cp = PureThy.get_thm thy2 "dj_cp"; + val pt_perm_compose = PureThy.get_thm thy2 "pt_perm_compose"; + val pt_perm_compose_rev = PureThy.get_thm thy2 "pt_perm_compose_rev"; + val dj_perm_perm_forget = PureThy.get_thm thy2 "dj_perm_perm_forget"; + + fun composition_instance name1 name2 thy = + let + val cp_class = cp_class_of thy name1 name2; + val pt_class = + if name1 = name2 then [pt_class_of thy name1] + else []; + val permT1 = mk_permT (Type (name1, [])); + val permT2 = mk_permT (Type (name2, [])); + val Ts = map body_type perm_types; + val cp_inst = cp_inst_of thy name1 name2; + val simps = simpset_of thy addsimps (perm_fun_def :: + (if name1 <> name2 then + let val dj = dj_thm_of thy name2 name1 + in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end + else + let + val at_inst = at_inst_of thy name1; + val pt_inst = pt_inst_of thy name1; + in + [cp_inst RS cp1 RS sym, + at_inst RS (pt_inst RS pt_perm_compose) RS sym, + at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym] + end)) + val sort = Sign.certify_sort thy (cp_class :: pt_class); + val thms = split_conj_thm (Goal.prove_global thy [] [] + (augment_sort thy sort + (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj + (map (fn ((s, T), x) => + let + val pi1 = Free ("pi1", permT1); + val pi2 = Free ("pi2", permT2); + val perm1 = Const (s, permT1 --> T --> T); + val perm2 = Const (s, permT2 --> T --> T); + val perm3 = Const ("Nominal.perm", permT1 --> permT2 --> permT2) + in HOLogic.mk_eq + (perm1 $ pi1 $ (perm2 $ pi2 $ Free (x, T)), + perm2 $ (perm3 $ pi1 $ pi2) $ (perm1 $ pi1 $ Free (x, T))) + end) + (perm_names ~~ Ts ~~ perm_indnames))))) + (fn _ => EVERY [indtac induction perm_indnames 1, + ALLGOALS (asm_full_simp_tac simps)])) + in + fold (fn (s, tvs) => fn thy => AxClass.prove_arity + (s, map (inter_sort thy sort o snd) tvs, [cp_class]) + (Class.intro_classes_tac [] THEN ALLGOALS (resolve_tac thms)) thy) + (full_new_type_names' ~~ tyvars) thy + end; + + val (perm_thmss,thy3) = thy2 |> + fold (fn name1 => fold (composition_instance name1) atoms) atoms |> + fold (fn atom => fn thy => + let val pt_name = pt_class_of thy atom + in + fold (fn (s, tvs) => fn thy => AxClass.prove_arity + (s, map (inter_sort thy [pt_name] o snd) tvs, [pt_name]) + (EVERY + [Class.intro_classes_tac [], + resolve_tac perm_empty_thms 1, + resolve_tac perm_append_thms 1, + resolve_tac perm_eq_thms 1, assume_tac 1]) thy) + (full_new_type_names' ~~ tyvars) thy + end) atoms |> + PureThy.add_thmss + [((Binding.name (space_implode "_" new_type_names ^ "_unfolded_perm_eq"), + unfolded_perm_eq_thms), [Simplifier.simp_add]), + ((Binding.name (space_implode "_" new_type_names ^ "_perm_empty"), + perm_empty_thms), [Simplifier.simp_add]), + ((Binding.name (space_implode "_" new_type_names ^ "_perm_append"), + perm_append_thms), [Simplifier.simp_add]), + ((Binding.name (space_implode "_" new_type_names ^ "_perm_eq"), + perm_eq_thms), [Simplifier.simp_add])]; + + (**** Define representing sets ****) + + val _ = warning "representing sets"; + + val rep_set_names = DatatypeProp.indexify_names + (map (fn (i, _) => name_of_typ (nth_dtyp i) ^ "_set") descr); + val big_rep_name = + space_implode "_" (DatatypeProp.indexify_names (List.mapPartial + (fn (i, ("Nominal.noption", _, _)) => NONE + | (i, _) => SOME (name_of_typ (nth_dtyp i))) descr)) ^ "_set"; + val _ = warning ("big_rep_name: " ^ big_rep_name); + + fun strip_option (dtf as DtType ("fun", [dt, DtRec i])) = + (case AList.lookup op = descr i of + SOME ("Nominal.noption", _, [(_, [dt']), _]) => + apfst (cons dt) (strip_option dt') + | _ => ([], dtf)) + | strip_option (DtType ("fun", [dt, DtType ("Nominal.noption", [dt'])])) = + apfst (cons dt) (strip_option dt') + | strip_option dt = ([], dt); + + val dt_atomTs = distinct op = (map (typ_of_dtyp descr sorts) + (List.concat (map (fn (_, (_, _, cs)) => List.concat + (map (List.concat o map (fst o strip_option) o snd) cs)) descr))); + val dt_atoms = map (fst o dest_Type) dt_atomTs; + + fun make_intr s T (cname, cargs) = + let + fun mk_prem (dt, (j, j', prems, ts)) = + let + val (dts, dt') = strip_option dt; + val (dts', dt'') = strip_dtyp dt'; + val Ts = map (typ_of_dtyp descr sorts) dts; + val Us = map (typ_of_dtyp descr sorts) dts'; + val T = typ_of_dtyp descr sorts dt''; + val free = mk_Free "x" (Us ---> T) j; + val free' = app_bnds free (length Us); + fun mk_abs_fun (T, (i, t)) = + let val U = fastype_of t + in (i + 1, Const ("Nominal.abs_fun", [T, U, T] ---> + Type ("Nominal.noption", [U])) $ mk_Free "y" T i $ t) + end + in (j + 1, j' + length Ts, + case dt'' of + DtRec k => list_all (map (pair "x") Us, + HOLogic.mk_Trueprop (Free (List.nth (rep_set_names, k), + T --> HOLogic.boolT) $ free')) :: prems + | _ => prems, + snd (List.foldr mk_abs_fun (j', free) Ts) :: ts) + end; + + val (_, _, prems, ts) = List.foldr mk_prem (1, 1, [], []) cargs; + val concl = HOLogic.mk_Trueprop (Free (s, T --> HOLogic.boolT) $ + list_comb (Const (cname, map fastype_of ts ---> T), ts)) + in Logic.list_implies (prems, concl) + end; + + val (intr_ts, (rep_set_names', recTs')) = + apfst List.concat (apsnd ListPair.unzip (ListPair.unzip (List.mapPartial + (fn ((_, ("Nominal.noption", _, _)), _) => NONE + | ((i, (_, _, constrs)), rep_set_name) => + let val T = nth_dtyp i + in SOME (map (make_intr rep_set_name T) constrs, + (rep_set_name, T)) + end) + (descr ~~ rep_set_names)))); + val rep_set_names'' = map (Sign.full_bname thy3) rep_set_names'; + + val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy4) = + Inductive.add_inductive_global (serial_string ()) + {quiet_mode = false, verbose = false, kind = Thm.internalK, + alt_name = Binding.name big_rep_name, coind = false, no_elim = true, no_ind = false, + skip_mono = true, fork_mono = false} + (map (fn (s, T) => ((Binding.name s, T --> HOLogic.boolT), NoSyn)) + (rep_set_names' ~~ recTs')) + [] (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] thy3; + + (**** Prove that representing set is closed under permutation ****) + + val _ = warning "proving closure under permutation..."; + + val abs_perm = PureThy.get_thms thy4 "abs_perm"; + + val perm_indnames' = List.mapPartial + (fn (x, (_, ("Nominal.noption", _, _))) => NONE | (x, _) => SOME x) + (perm_indnames ~~ descr); + + fun mk_perm_closed name = map (fn th => standard (th RS mp)) + (List.take (split_conj_thm (Goal.prove_global thy4 [] [] + (augment_sort thy4 + (pt_class_of thy4 name :: map (cp_class_of thy4 name) (dt_atoms \ name)) + (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map + (fn ((s, T), x) => + let + val S = Const (s, T --> HOLogic.boolT); + val permT = mk_permT (Type (name, [])) + in HOLogic.mk_imp (S $ Free (x, T), + S $ (Const ("Nominal.perm", permT --> T --> T) $ + Free ("pi", permT) $ Free (x, T))) + end) (rep_set_names'' ~~ recTs' ~~ perm_indnames'))))) + (fn _ => EVERY + [indtac rep_induct [] 1, + ALLGOALS (simp_tac (simpset_of thy4 addsimps + (symmetric perm_fun_def :: abs_perm))), + ALLGOALS (resolve_tac rep_intrs THEN_ALL_NEW assume_tac)])), + length new_type_names)); + + val perm_closed_thmss = map mk_perm_closed atoms; + + (**** typedef ****) + + val _ = warning "defining type..."; + + val (typedefs, thy6) = + thy4 + |> fold_map (fn ((((name, mx), tvs), (cname, U)), name') => fn thy => + Typedef.add_typedef false (SOME (Binding.name name')) + (Binding.name name, map fst tvs, mx) + (Const ("Collect", (U --> HOLogic.boolT) --> HOLogic.mk_setT U) $ + Const (cname, U --> HOLogic.boolT)) NONE + (rtac exI 1 THEN rtac CollectI 1 THEN + QUIET_BREADTH_FIRST (has_fewer_prems 1) + (resolve_tac rep_intrs 1)) thy |> (fn ((_, r), thy) => + let + val permT = mk_permT + (TFree (Name.variant (map fst tvs) "'a", HOLogic.typeS)); + val pi = Free ("pi", permT); + val T = Type (Sign.intern_type thy name, map TFree tvs); + in apfst (pair r o hd) + (PureThy.add_defs_unchecked true [((Binding.name ("prm_" ^ name ^ "_def"), Logic.mk_equals + (Const ("Nominal.perm", permT --> T --> T) $ pi $ Free ("x", T), + Const (Sign.intern_const thy ("Abs_" ^ name), U --> T) $ + (Const ("Nominal.perm", permT --> U --> U) $ pi $ + (Const (Sign.intern_const thy ("Rep_" ^ name), T --> U) $ + Free ("x", T))))), [])] thy) + end)) + (types_syntax ~~ tyvars ~~ + List.take (rep_set_names'' ~~ recTs', length new_type_names) ~~ + new_type_names); + + val perm_defs = map snd typedefs; + val Abs_inverse_thms = map (collect_simp o #Abs_inverse o fst) typedefs; + val Rep_inverse_thms = map (#Rep_inverse o fst) typedefs; + val Rep_thms = map (collect_simp o #Rep o fst) typedefs; + + + (** prove that new types are in class pt_ **) + + val _ = warning "prove that new types are in class pt_ ..."; + + fun pt_instance (atom, perm_closed_thms) = + fold (fn ((((((Abs_inverse, Rep_inverse), Rep), + perm_def), name), tvs), perm_closed) => fn thy => + let + val pt_class = pt_class_of thy atom; + val sort = Sign.certify_sort thy + (pt_class :: map (cp_class_of thy atom) (dt_atoms \ atom)) + in AxClass.prove_arity + (Sign.intern_type thy name, + map (inter_sort thy sort o snd) tvs, [pt_class]) + (EVERY [Class.intro_classes_tac [], + rewrite_goals_tac [perm_def], + asm_full_simp_tac (simpset_of thy addsimps [Rep_inverse]) 1, + asm_full_simp_tac (simpset_of thy addsimps + [Rep RS perm_closed RS Abs_inverse]) 1, + asm_full_simp_tac (HOL_basic_ss addsimps [PureThy.get_thm thy + ("pt_" ^ Long_Name.base_name atom ^ "3")]) 1]) thy + end) + (Abs_inverse_thms ~~ Rep_inverse_thms ~~ Rep_thms ~~ perm_defs ~~ + new_type_names ~~ tyvars ~~ perm_closed_thms); + + + (** prove that new types are in class cp__ **) + + val _ = warning "prove that new types are in class cp__ ..."; + + fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy = + let + val cp_class = cp_class_of thy atom1 atom2; + val sort = Sign.certify_sort thy + (pt_class_of thy atom1 :: map (cp_class_of thy atom1) (dt_atoms \ atom1) @ + (if atom1 = atom2 then [cp_class_of thy atom1 atom1] else + pt_class_of thy atom2 :: map (cp_class_of thy atom2) (dt_atoms \ atom2))); + val cp1' = cp_inst_of thy atom1 atom2 RS cp1 + in fold (fn ((((((Abs_inverse, Rep), + perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy => + AxClass.prove_arity + (Sign.intern_type thy name, + map (inter_sort thy sort o snd) tvs, [cp_class]) + (EVERY [Class.intro_classes_tac [], + rewrite_goals_tac [perm_def], + asm_full_simp_tac (simpset_of thy addsimps + ((Rep RS perm_closed1 RS Abs_inverse) :: + (if atom1 = atom2 then [] + else [Rep RS perm_closed2 RS Abs_inverse]))) 1, + cong_tac 1, + rtac refl 1, + rtac cp1' 1]) thy) + (Abs_inverse_thms ~~ Rep_thms ~~ perm_defs ~~ new_type_names ~~ + tyvars ~~ perm_closed_thms1 ~~ perm_closed_thms2) thy + end; + + val thy7 = fold (fn x => fn thy => thy |> + pt_instance x |> + fold (cp_instance x) (atoms ~~ perm_closed_thmss)) + (atoms ~~ perm_closed_thmss) thy6; + + (**** constructors ****) + + fun mk_abs_fun (x, t) = + let + val T = fastype_of x; + val U = fastype_of t + in + Const ("Nominal.abs_fun", T --> U --> T --> + Type ("Nominal.noption", [U])) $ x $ t + end; + + val (ty_idxs, _) = List.foldl + (fn ((i, ("Nominal.noption", _, _)), p) => p + | ((i, _), (ty_idxs, j)) => (ty_idxs @ [(i, j)], j + 1)) ([], 0) descr; + + fun reindex (DtType (s, dts)) = DtType (s, map reindex dts) + | reindex (DtRec i) = DtRec (the (AList.lookup op = ty_idxs i)) + | reindex dt = dt; + + fun strip_suffix i s = implode (List.take (explode s, size s - i)); + + (** strips the "_Rep" in type names *) + fun strip_nth_name i s = + let val xs = Long_Name.explode s; + in Long_Name.implode (Library.nth_map (length xs - i) (strip_suffix 4) xs) end; + + val (descr'', ndescr) = ListPair.unzip (map_filter + (fn (i, ("Nominal.noption", _, _)) => NONE + | (i, (s, dts, constrs)) => + let + val SOME index = AList.lookup op = ty_idxs i; + val (constrs2, constrs1) = + map_split (fn (cname, cargs) => + apsnd (pair (strip_nth_name 2 (strip_nth_name 1 cname))) + (fold_map (fn dt => fn dts => + let val (dts', dt') = strip_option dt + in ((length dts, length dts'), dts @ dts' @ [reindex dt']) end) + cargs [])) constrs + in SOME ((index, (strip_nth_name 1 s, map reindex dts, constrs1)), + (index, constrs2)) + end) descr); + + val (descr1, descr2) = chop (length new_type_names) descr''; + val descr' = [descr1, descr2]; + + fun partition_cargs idxs xs = map (fn (i, j) => + (List.take (List.drop (xs, i), j), List.nth (xs, i + j))) idxs; + + val pdescr = map (fn ((i, (s, dts, constrs)), (_, idxss)) => (i, (s, dts, + map (fn ((cname, cargs), idxs) => (cname, partition_cargs idxs cargs)) + (constrs ~~ idxss)))) (descr'' ~~ ndescr); + + fun nth_dtyp' i = typ_of_dtyp descr'' sorts (DtRec i); + + val rep_names = map (fn s => + Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names; + val abs_names = map (fn s => + Sign.intern_const thy7 ("Abs_" ^ s)) new_type_names; + + val recTs = get_rec_types descr'' sorts; + val newTs' = Library.take (length new_type_names, recTs'); + val newTs = Library.take (length new_type_names, recTs); + + val full_new_type_names = map (Sign.full_bname thy) new_type_names; + + fun make_constr_def tname T T' ((thy, defs, eqns), + (((cname_rep, _), (cname, cargs)), (cname', mx))) = + let + fun constr_arg ((dts, dt), (j, l_args, r_args)) = + let + val xs = map (fn (dt, i) => mk_Free "x" (typ_of_dtyp descr'' sorts dt) i) + (dts ~~ (j upto j + length dts - 1)) + val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts) + in + (j + length dts + 1, + xs @ x :: l_args, + List.foldr mk_abs_fun + (case dt of + DtRec k => if k < length new_type_names then + Const (List.nth (rep_names, k), typ_of_dtyp descr'' sorts dt --> + typ_of_dtyp descr sorts dt) $ x + else error "nested recursion not (yet) supported" + | _ => x) xs :: r_args) + end + + val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) cargs; + val abs_name = Sign.intern_const thy ("Abs_" ^ tname); + val rep_name = Sign.intern_const thy ("Rep_" ^ tname); + val constrT = map fastype_of l_args ---> T; + val lhs = list_comb (Const (cname, constrT), l_args); + val rhs = list_comb (Const (cname_rep, map fastype_of r_args ---> T'), r_args); + val def = Logic.mk_equals (lhs, Const (abs_name, T' --> T) $ rhs); + val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq + (Const (rep_name, T --> T') $ lhs, rhs)); + val def_name = (Long_Name.base_name cname) ^ "_def"; + val ([def_thm], thy') = thy |> + Sign.add_consts_i [(Binding.name cname', constrT, mx)] |> + (PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)] + in (thy', defs @ [def_thm], eqns @ [eqn]) end; + + fun dt_constr_defs ((thy, defs, eqns, dist_lemmas), ((((((_, (_, _, constrs)), + (_, (_, _, constrs'))), tname), T), T'), constr_syntax)) = + let + val rep_const = cterm_of thy + (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T')); + val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma); + val (thy', defs', eqns') = Library.foldl (make_constr_def tname T T') + ((Sign.add_path tname thy, defs, []), constrs ~~ constrs' ~~ constr_syntax) + in + (parent_path (#flat_names config) thy', defs', eqns @ [eqns'], dist_lemmas @ [dist]) + end; + + val (thy8, constr_defs, constr_rep_eqns, dist_lemmas) = Library.foldl dt_constr_defs + ((thy7, [], [], []), List.take (descr, length new_type_names) ~~ + List.take (pdescr, length new_type_names) ~~ + new_type_names ~~ newTs ~~ newTs' ~~ constr_syntax); + + val abs_inject_thms = map (collect_simp o #Abs_inject o fst) typedefs + val rep_inject_thms = map (#Rep_inject o fst) typedefs + + (* prove theorem Rep_i (Constr_j ...) = Constr'_j ... *) + + fun prove_constr_rep_thm eqn = + let + val inj_thms = map (fn r => r RS iffD1) abs_inject_thms; + val rewrites = constr_defs @ map mk_meta_eq Rep_inverse_thms + in Goal.prove_global thy8 [] [] eqn (fn _ => EVERY + [resolve_tac inj_thms 1, + rewrite_goals_tac rewrites, + rtac refl 3, + resolve_tac rep_intrs 2, + REPEAT (resolve_tac Rep_thms 1)]) + end; + + val constr_rep_thmss = map (map prove_constr_rep_thm) constr_rep_eqns; + + (* prove theorem pi \ Rep_i x = Rep_i (pi \ x) *) + + fun prove_perm_rep_perm (atom, perm_closed_thms) = map (fn th => + let + val _ $ (_ $ (Rep $ x)) = Logic.unvarify (prop_of th); + val Type ("fun", [T, U]) = fastype_of Rep; + val permT = mk_permT (Type (atom, [])); + val pi = Free ("pi", permT); + in + Goal.prove_global thy8 [] [] + (augment_sort thy8 + (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (dt_atoms \ atom)) + (HOLogic.mk_Trueprop (HOLogic.mk_eq + (Const ("Nominal.perm", permT --> U --> U) $ pi $ (Rep $ x), + Rep $ (Const ("Nominal.perm", permT --> T --> T) $ pi $ x))))) + (fn _ => simp_tac (HOL_basic_ss addsimps (perm_defs @ Abs_inverse_thms @ + perm_closed_thms @ Rep_thms)) 1) + end) Rep_thms; + + val perm_rep_perm_thms = List.concat (map prove_perm_rep_perm + (atoms ~~ perm_closed_thmss)); + + (* prove distinctness theorems *) + + val distinct_props = DatatypeProp.make_distincts descr' sorts; + val dist_rewrites = map2 (fn rep_thms => fn dist_lemma => + dist_lemma :: rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0]) + constr_rep_thmss dist_lemmas; + + fun prove_distinct_thms _ (_, []) = [] + | prove_distinct_thms (p as (rep_thms, dist_lemma)) (k, t :: ts) = + let + val dist_thm = Goal.prove_global thy8 [] [] t (fn _ => + simp_tac (simpset_of thy8 addsimps (dist_lemma :: rep_thms)) 1) + in dist_thm :: standard (dist_thm RS not_sym) :: + prove_distinct_thms p (k, ts) + end; + + val distinct_thms = map2 prove_distinct_thms + (constr_rep_thmss ~~ dist_lemmas) distinct_props; + + (** prove equations for permutation functions **) + + val perm_simps' = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) => + let val T = nth_dtyp' i + in List.concat (map (fn (atom, perm_closed_thms) => + map (fn ((cname, dts), constr_rep_thm) => + let + val cname = Sign.intern_const thy8 + (Long_Name.append tname (Long_Name.base_name cname)); + val permT = mk_permT (Type (atom, [])); + val pi = Free ("pi", permT); + + fun perm t = + let val T = fastype_of t + in Const ("Nominal.perm", permT --> T --> T) $ pi $ t end; + + fun constr_arg ((dts, dt), (j, l_args, r_args)) = + let + val Ts = map (typ_of_dtyp descr'' sorts) dts; + val xs = map (fn (T, i) => mk_Free "x" T i) + (Ts ~~ (j upto j + length dts - 1)) + val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts) + in + (j + length dts + 1, + xs @ x :: l_args, + map perm (xs @ [x]) @ r_args) + end + + val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) dts; + val c = Const (cname, map fastype_of l_args ---> T) + in + Goal.prove_global thy8 [] [] + (augment_sort thy8 + (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (dt_atoms \ atom)) + (HOLogic.mk_Trueprop (HOLogic.mk_eq + (perm (list_comb (c, l_args)), list_comb (c, r_args))))) + (fn _ => EVERY + [simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: perm_defs)) 1, + simp_tac (HOL_basic_ss addsimps (Rep_thms @ Abs_inverse_thms @ + constr_defs @ perm_closed_thms)) 1, + TRY (simp_tac (HOL_basic_ss addsimps + (symmetric perm_fun_def :: abs_perm)) 1), + TRY (simp_tac (HOL_basic_ss addsimps + (perm_fun_def :: perm_defs @ Rep_thms @ Abs_inverse_thms @ + perm_closed_thms)) 1)]) + end) (constrs ~~ constr_rep_thms)) (atoms ~~ perm_closed_thmss)) + end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss); + + (** prove injectivity of constructors **) + + val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms; + val alpha = PureThy.get_thms thy8 "alpha"; + val abs_fresh = PureThy.get_thms thy8 "abs_fresh"; + + val pt_cp_sort = + map (pt_class_of thy8) dt_atoms @ + maps (fn s => map (cp_class_of thy8 s) (dt_atoms \ s)) dt_atoms; + + val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) => + let val T = nth_dtyp' i + in List.mapPartial (fn ((cname, dts), constr_rep_thm) => + if null dts then NONE else SOME + let + val cname = Sign.intern_const thy8 + (Long_Name.append tname (Long_Name.base_name cname)); + + fun make_inj ((dts, dt), (j, args1, args2, eqs)) = + let + val Ts_idx = map (typ_of_dtyp descr'' sorts) dts ~~ (j upto j + length dts - 1); + val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx; + val ys = map (fn (T, i) => mk_Free "y" T i) Ts_idx; + val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts); + val y = mk_Free "y" (typ_of_dtyp descr'' sorts dt) (j + length dts) + in + (j + length dts + 1, + xs @ (x :: args1), ys @ (y :: args2), + HOLogic.mk_eq + (List.foldr mk_abs_fun x xs, List.foldr mk_abs_fun y ys) :: eqs) + end; + + val (_, args1, args2, eqs) = List.foldr make_inj (1, [], [], []) dts; + val Ts = map fastype_of args1; + val c = Const (cname, Ts ---> T) + in + Goal.prove_global thy8 [] [] + (augment_sort thy8 pt_cp_sort + (HOLogic.mk_Trueprop (HOLogic.mk_eq + (HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)), + foldr1 HOLogic.mk_conj eqs)))) + (fn _ => EVERY + [asm_full_simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: + rep_inject_thms')) 1, + TRY (asm_full_simp_tac (HOL_basic_ss addsimps (fresh_def :: supp_def :: + alpha @ abs_perm @ abs_fresh @ rep_inject_thms @ + perm_rep_perm_thms)) 1)]) + end) (constrs ~~ constr_rep_thms) + end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss); + + (** equations for support and freshness **) + + val (supp_thms, fresh_thms) = ListPair.unzip (map ListPair.unzip + (map (fn ((((i, (_, _, constrs)), tname), inject_thms'), perm_thms') => + let val T = nth_dtyp' i + in List.concat (map (fn (cname, dts) => map (fn atom => + let + val cname = Sign.intern_const thy8 + (Long_Name.append tname (Long_Name.base_name cname)); + val atomT = Type (atom, []); + + fun process_constr ((dts, dt), (j, args1, args2)) = + let + val Ts_idx = map (typ_of_dtyp descr'' sorts) dts ~~ (j upto j + length dts - 1); + val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx; + val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts) + in + (j + length dts + 1, + xs @ (x :: args1), List.foldr mk_abs_fun x xs :: args2) + end; + + val (_, args1, args2) = List.foldr process_constr (1, [], []) dts; + val Ts = map fastype_of args1; + val c = list_comb (Const (cname, Ts ---> T), args1); + fun supp t = + Const ("Nominal.supp", fastype_of t --> HOLogic.mk_setT atomT) $ t; + fun fresh t = fresh_const atomT (fastype_of t) $ Free ("a", atomT) $ t; + val supp_thm = Goal.prove_global thy8 [] [] + (augment_sort thy8 pt_cp_sort + (HOLogic.mk_Trueprop (HOLogic.mk_eq + (supp c, + if null dts then HOLogic.mk_set atomT [] + else foldr1 (HOLogic.mk_binop @{const_name Un}) (map supp args2))))) + (fn _ => + simp_tac (HOL_basic_ss addsimps (supp_def :: + Un_assoc :: de_Morgan_conj :: Collect_disj_eq :: finite_Un :: + symmetric empty_def :: finite_emptyI :: simp_thms @ + abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1) + in + (supp_thm, + Goal.prove_global thy8 [] [] (augment_sort thy8 pt_cp_sort + (HOLogic.mk_Trueprop (HOLogic.mk_eq + (fresh c, + if null dts then HOLogic.true_const + else foldr1 HOLogic.mk_conj (map fresh args2))))) + (fn _ => + simp_tac (HOL_ss addsimps [Un_iff, empty_iff, fresh_def, supp_thm]) 1)) + end) atoms) constrs) + end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ inject_thms ~~ perm_simps'))); + + (**** weak induction theorem ****) + + fun mk_indrule_lemma ((prems, concls), (((i, _), T), U)) = + let + val Rep_t = Const (List.nth (rep_names, i), T --> U) $ + mk_Free "x" T i; + + val Abs_t = Const (List.nth (abs_names, i), U --> T) + + in (prems @ [HOLogic.imp $ + (Const (List.nth (rep_set_names'', i), U --> HOLogic.boolT) $ Rep_t) $ + (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))], + concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i]) + end; + + val (indrule_lemma_prems, indrule_lemma_concls) = + Library.foldl mk_indrule_lemma (([], []), (descr'' ~~ recTs ~~ recTs')); + + val indrule_lemma = Goal.prove_global thy8 [] [] + (Logic.mk_implies + (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems), + HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY + [REPEAT (etac conjE 1), + REPEAT (EVERY + [TRY (rtac conjI 1), full_simp_tac (HOL_basic_ss addsimps Rep_inverse_thms) 1, + etac mp 1, resolve_tac Rep_thms 1])]); + + val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma))); + val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else + map (Free o apfst fst o dest_Var) Ps; + val indrule_lemma' = cterm_instantiate + (map (cterm_of thy8) Ps ~~ map (cterm_of thy8) frees) indrule_lemma; + + val Abs_inverse_thms' = map (fn r => r RS subst) Abs_inverse_thms; + + val dt_induct_prop = DatatypeProp.make_ind descr' sorts; + val dt_induct = Goal.prove_global thy8 [] + (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop) + (fn {prems, ...} => EVERY + [rtac indrule_lemma' 1, + (indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, + EVERY (map (fn (prem, r) => (EVERY + [REPEAT (eresolve_tac Abs_inverse_thms' 1), + simp_tac (HOL_basic_ss addsimps [symmetric r]) 1, + DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)])) + (prems ~~ constr_defs))]); + + val case_names_induct = mk_case_names_induct descr''; + + (**** prove that new datatypes have finite support ****) + + val _ = warning "proving finite support for the new datatype"; + + val indnames = DatatypeProp.make_tnames recTs; + + val abs_supp = PureThy.get_thms thy8 "abs_supp"; + val supp_atm = PureThy.get_thms thy8 "supp_atm"; + + val finite_supp_thms = map (fn atom => + let val atomT = Type (atom, []) + in map standard (List.take + (split_conj_thm (Goal.prove_global thy8 [] [] + (augment_sort thy8 (fs_class_of thy8 atom :: pt_cp_sort) + (HOLogic.mk_Trueprop + (foldr1 HOLogic.mk_conj (map (fn (s, T) => + Const ("Finite_Set.finite", HOLogic.mk_setT atomT --> HOLogic.boolT) $ + (Const ("Nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T))) + (indnames ~~ recTs))))) + (fn _ => indtac dt_induct indnames 1 THEN + ALLGOALS (asm_full_simp_tac (simpset_of thy8 addsimps + (abs_supp @ supp_atm @ + PureThy.get_thms thy8 ("fs_" ^ Long_Name.base_name atom ^ "1") @ + List.concat supp_thms))))), + length new_type_names)) + end) atoms; + + val simp_atts = replicate (length new_type_names) [Simplifier.simp_add]; + + (* Function to add both the simp and eqvt attributes *) + (* These two attributes are duplicated on all the types in the mutual nominal datatypes *) + + val simp_eqvt_atts = replicate (length new_type_names) [Simplifier.simp_add, NominalThmDecls.eqvt_add]; + + val (_, thy9) = thy8 |> + Sign.add_path big_name |> + PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])] ||>> + PureThy.add_thmss [((Binding.name "inducts", projections dt_induct), [case_names_induct])] ||> + Sign.parent_path ||>> + DatatypeAux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>> + DatatypeAux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>> + DatatypeAux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>> + DatatypeAux.store_thmss "inject" new_type_names inject_thms ||>> + DatatypeAux.store_thmss "supp" new_type_names supp_thms ||>> + DatatypeAux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||> + fold (fn (atom, ths) => fn thy => + let + val class = fs_class_of thy atom; + val sort = Sign.certify_sort thy (class :: pt_cp_sort) + in fold (fn Type (s, Ts) => AxClass.prove_arity + (s, map (inter_sort thy sort o snd o dest_TFree) Ts, [class]) + (Class.intro_classes_tac [] THEN resolve_tac ths 1)) newTs thy + end) (atoms ~~ finite_supp_thms); + + (**** strong induction theorem ****) + + val pnames = if length descr'' = 1 then ["P"] + else map (fn i => "P" ^ string_of_int i) (1 upto length descr''); + val ind_sort = if null dt_atomTs then HOLogic.typeS + else Sign.certify_sort thy9 (map (fs_class_of thy9) dt_atoms); + val fsT = TFree ("'n", ind_sort); + val fsT' = TFree ("'n", HOLogic.typeS); + + val fresh_fs = map (fn (s, T) => (T, Free (s, fsT' --> HOLogic.mk_setT T))) + (DatatypeProp.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs); + + fun make_pred fsT i T = + Free (List.nth (pnames, i), fsT --> T --> HOLogic.boolT); + + fun mk_fresh1 xs [] = [] + | mk_fresh1 xs ((y as (_, T)) :: ys) = map (fn x => HOLogic.mk_Trueprop + (HOLogic.mk_not (HOLogic.mk_eq (Free y, Free x)))) + (filter (fn (_, U) => T = U) (rev xs)) @ + mk_fresh1 (y :: xs) ys; + + fun mk_fresh2 xss [] = [] + | mk_fresh2 xss ((p as (ys, _)) :: yss) = List.concat (map (fn y as (_, T) => + map (fn (_, x as (_, U)) => HOLogic.mk_Trueprop + (fresh_const T U $ Free y $ Free x)) (rev xss @ yss)) ys) @ + mk_fresh2 (p :: xss) yss; + + fun make_ind_prem fsT f k T ((cname, cargs), idxs) = + let + val recs = List.filter is_rec_type cargs; + val Ts = map (typ_of_dtyp descr'' sorts) cargs; + val recTs' = map (typ_of_dtyp descr'' sorts) recs; + val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts); + val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs)); + val frees = tnames ~~ Ts; + val frees' = partition_cargs idxs frees; + val z = (Name.variant tnames "z", fsT); + + fun mk_prem ((dt, s), T) = + let + val (Us, U) = strip_type T; + val l = length Us + in list_all (z :: map (pair "x") Us, HOLogic.mk_Trueprop + (make_pred fsT (body_index dt) U $ Bound l $ app_bnds (Free (s, T)) l)) + end; + + val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs'); + val prems' = map (fn p as (_, T) => HOLogic.mk_Trueprop + (f T (Free p) (Free z))) (List.concat (map fst frees')) @ + mk_fresh1 [] (List.concat (map fst frees')) @ + mk_fresh2 [] frees' + + in list_all_free (frees @ [z], Logic.list_implies (prems' @ prems, + HOLogic.mk_Trueprop (make_pred fsT k T $ Free z $ + list_comb (Const (cname, Ts ---> T), map Free frees)))) + end; + + val ind_prems = List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) => + map (make_ind_prem fsT (fn T => fn t => fn u => + fresh_const T fsT $ t $ u) i T) + (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs)); + val tnames = DatatypeProp.make_tnames recTs; + val zs = Name.variant_list tnames (replicate (length descr'') "z"); + val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &") + (map (fn ((((i, _), T), tname), z) => + make_pred fsT i T $ Free (z, fsT) $ Free (tname, T)) + (descr'' ~~ recTs ~~ tnames ~~ zs))); + val induct = Logic.list_implies (ind_prems, ind_concl); + + val ind_prems' = + map (fn (_, f as Free (_, T)) => list_all_free ([("x", fsT')], + HOLogic.mk_Trueprop (Const ("Finite_Set.finite", + (snd (split_last (binder_types T)) --> HOLogic.boolT) --> + HOLogic.boolT) $ (f $ Free ("x", fsT'))))) fresh_fs @ + List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) => + map (make_ind_prem fsT' (fn T => fn t => fn u => HOLogic.Not $ + HOLogic.mk_mem (t, the (AList.lookup op = fresh_fs T) $ u)) i T) + (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs)); + val ind_concl' = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &") + (map (fn ((((i, _), T), tname), z) => + make_pred fsT' i T $ Free (z, fsT') $ Free (tname, T)) + (descr'' ~~ recTs ~~ tnames ~~ zs))); + val induct' = Logic.list_implies (ind_prems', ind_concl'); + + val aux_ind_vars = + (DatatypeProp.indexify_names (replicate (length dt_atomTs) "pi") ~~ + map mk_permT dt_atomTs) @ [("z", fsT')]; + val aux_ind_Ts = rev (map snd aux_ind_vars); + val aux_ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &") + (map (fn (((i, _), T), tname) => + HOLogic.list_all (aux_ind_vars, make_pred fsT' i T $ Bound 0 $ + fold_rev (mk_perm aux_ind_Ts) (map Bound (length dt_atomTs downto 1)) + (Free (tname, T)))) + (descr'' ~~ recTs ~~ tnames))); + + val fin_set_supp = map (fn s => + at_inst_of thy9 s RS at_fin_set_supp) dt_atoms; + val fin_set_fresh = map (fn s => + at_inst_of thy9 s RS at_fin_set_fresh) dt_atoms; + val pt1_atoms = map (fn Type (s, _) => + PureThy.get_thm thy9 ("pt_" ^ Long_Name.base_name s ^ "1")) dt_atomTs; + val pt2_atoms = map (fn Type (s, _) => + PureThy.get_thm thy9 ("pt_" ^ Long_Name.base_name s ^ "2") RS sym) dt_atomTs; + val exists_fresh' = PureThy.get_thms thy9 "exists_fresh'"; + val fs_atoms = PureThy.get_thms thy9 "fin_supp"; + val abs_supp = PureThy.get_thms thy9 "abs_supp"; + val perm_fresh_fresh = PureThy.get_thms thy9 "perm_fresh_fresh"; + val calc_atm = PureThy.get_thms thy9 "calc_atm"; + val fresh_atm = PureThy.get_thms thy9 "fresh_atm"; + val fresh_left = PureThy.get_thms thy9 "fresh_left"; + val perm_swap = PureThy.get_thms thy9 "perm_swap"; + + fun obtain_fresh_name' ths ts T (freshs1, freshs2, ctxt) = + let + val p = foldr1 HOLogic.mk_prod (ts @ freshs1); + val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop + (HOLogic.exists_const T $ Abs ("x", T, + fresh_const T (fastype_of p) $ + Bound 0 $ p))) + (fn _ => EVERY + [resolve_tac exists_fresh' 1, + simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms @ + fin_set_supp @ ths)) 1]); + val (([cx], ths), ctxt') = Obtain.result + (fn _ => EVERY + [etac exE 1, + full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1, + REPEAT (etac conjE 1)]) + [ex] ctxt + in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end; + + fun fresh_fresh_inst thy a b = + let + val T = fastype_of a; + val SOME th = find_first (fn th => case prop_of th of + _ $ (_ $ (Const (_, Type (_, [U, _])) $ _ $ _)) $ _ => U = T + | _ => false) perm_fresh_fresh + in + Drule.instantiate' [] + [SOME (cterm_of thy a), NONE, SOME (cterm_of thy b)] th + end; + + val fs_cp_sort = + map (fs_class_of thy9) dt_atoms @ + maps (fn s => map (cp_class_of thy9 s) (dt_atoms \ s)) dt_atoms; + + (********************************************************************** + The subgoals occurring in the proof of induct_aux have the + following parameters: + + x_1 ... x_k p_1 ... p_m z + + where + + x_i : constructor arguments (introduced by weak induction rule) + p_i : permutations (one for each atom type in the data type) + z : freshness context + ***********************************************************************) + + val _ = warning "proving strong induction theorem ..."; + + val induct_aux = Goal.prove_global thy9 [] + (map (augment_sort thy9 fs_cp_sort) ind_prems') + (augment_sort thy9 fs_cp_sort ind_concl') (fn {prems, context} => + let + val (prems1, prems2) = chop (length dt_atomTs) prems; + val ind_ss2 = HOL_ss addsimps + finite_Diff :: abs_fresh @ abs_supp @ fs_atoms; + val ind_ss1 = ind_ss2 addsimps fresh_left @ calc_atm @ + fresh_atm @ rev_simps @ app_simps; + val ind_ss3 = HOL_ss addsimps abs_fun_eq1 :: + abs_perm @ calc_atm @ perm_swap; + val ind_ss4 = HOL_basic_ss addsimps fresh_left @ prems1 @ + fin_set_fresh @ calc_atm; + val ind_ss5 = HOL_basic_ss addsimps pt1_atoms; + val ind_ss6 = HOL_basic_ss addsimps flat perm_simps'; + val th = Goal.prove context [] [] + (augment_sort thy9 fs_cp_sort aux_ind_concl) + (fn {context = context1, ...} => + EVERY (indtac dt_induct tnames 1 :: + maps (fn ((_, (_, _, constrs)), (_, constrs')) => + map (fn ((cname, cargs), is) => + REPEAT (rtac allI 1) THEN + SUBPROOF (fn {prems = iprems, params, concl, + context = context2, ...} => + let + val concl' = term_of concl; + val _ $ (_ $ _ $ u) = concl'; + val U = fastype_of u; + val (xs, params') = + chop (length cargs) (map term_of params); + val Ts = map fastype_of xs; + val cnstr = Const (cname, Ts ---> U); + val (pis, z) = split_last params'; + val mk_pi = fold_rev (mk_perm []) pis; + val xs' = partition_cargs is xs; + val xs'' = map (fn (ts, u) => (map mk_pi ts, mk_pi u)) xs'; + val ts = maps (fn (ts, u) => ts @ [u]) xs''; + val (freshs1, freshs2, context3) = fold (fn t => + let val T = fastype_of t + in obtain_fresh_name' prems1 + (the (AList.lookup op = fresh_fs T) $ z :: ts) T + end) (maps fst xs') ([], [], context2); + val freshs1' = unflat (map fst xs') freshs1; + val freshs2' = map (Simplifier.simplify ind_ss4) + (mk_not_sym freshs2); + val ind_ss1' = ind_ss1 addsimps freshs2'; + val ind_ss3' = ind_ss3 addsimps freshs2'; + val rename_eq = + if forall (null o fst) xs' then [] + else [Goal.prove context3 [] [] + (HOLogic.mk_Trueprop (HOLogic.mk_eq + (list_comb (cnstr, ts), + list_comb (cnstr, maps (fn ((bs, t), cs) => + cs @ [fold_rev (mk_perm []) (map perm_of_pair + (bs ~~ cs)) t]) (xs'' ~~ freshs1'))))) + (fn _ => EVERY + (simp_tac (HOL_ss addsimps flat inject_thms) 1 :: + REPEAT (FIRSTGOAL (rtac conjI)) :: + maps (fn ((bs, t), cs) => + if null bs then [] + else rtac sym 1 :: maps (fn (b, c) => + [rtac trans 1, rtac sym 1, + rtac (fresh_fresh_inst thy9 b c) 1, + simp_tac ind_ss1' 1, + simp_tac ind_ss2 1, + simp_tac ind_ss3' 1]) (bs ~~ cs)) + (xs'' ~~ freshs1')))]; + val th = Goal.prove context3 [] [] concl' (fn _ => EVERY + [simp_tac (ind_ss6 addsimps rename_eq) 1, + cut_facts_tac iprems 1, + (resolve_tac prems THEN_ALL_NEW + SUBGOAL (fn (t, i) => case Logic.strip_assums_concl t of + _ $ (Const ("Nominal.fresh", _) $ _ $ _) => + simp_tac ind_ss1' i + | _ $ (Const ("Not", _) $ _) => + resolve_tac freshs2' i + | _ => asm_simp_tac (HOL_basic_ss addsimps + pt2_atoms addsimprocs [perm_simproc]) i)) 1]) + val final = ProofContext.export context3 context2 [th] + in + resolve_tac final 1 + end) context1 1) (constrs ~~ constrs')) (descr'' ~~ ndescr))) + in + EVERY + [cut_facts_tac [th] 1, + REPEAT (eresolve_tac [conjE, @{thm allE_Nil}] 1), + REPEAT (etac allE 1), + REPEAT (TRY (rtac conjI 1) THEN asm_full_simp_tac ind_ss5 1)] + end); + + val induct_aux' = Thm.instantiate ([], + map (fn (s, v as Var (_, T)) => + (cterm_of thy9 v, cterm_of thy9 (Free (s, T)))) + (pnames ~~ map head_of (HOLogic.dest_conj + (HOLogic.dest_Trueprop (concl_of induct_aux)))) @ + map (fn (_, f) => + let val f' = Logic.varify f + in (cterm_of thy9 f', + cterm_of thy9 (Const ("Nominal.supp", fastype_of f'))) + end) fresh_fs) induct_aux; + + val induct = Goal.prove_global thy9 [] + (map (augment_sort thy9 fs_cp_sort) ind_prems) + (augment_sort thy9 fs_cp_sort ind_concl) + (fn {prems, ...} => EVERY + [rtac induct_aux' 1, + REPEAT (resolve_tac fs_atoms 1), + REPEAT ((resolve_tac prems THEN_ALL_NEW + (etac meta_spec ORELSE' full_simp_tac (HOL_basic_ss addsimps [fresh_def]))) 1)]) + + val (_, thy10) = thy9 |> + Sign.add_path big_name |> + PureThy.add_thms [((Binding.name "strong_induct'", induct_aux), [])] ||>> + PureThy.add_thms [((Binding.name "strong_induct", induct), [case_names_induct])] ||>> + PureThy.add_thmss [((Binding.name "strong_inducts", projections induct), [case_names_induct])]; + + (**** recursion combinator ****) + + val _ = warning "defining recursion combinator ..."; + + val used = List.foldr OldTerm.add_typ_tfree_names [] recTs; + + val (rec_result_Ts', rec_fn_Ts') = DatatypeProp.make_primrec_Ts descr' sorts used; + + val rec_sort = if null dt_atomTs then HOLogic.typeS else + Sign.certify_sort thy10 pt_cp_sort; + + val rec_result_Ts = map (fn TFree (s, _) => TFree (s, rec_sort)) rec_result_Ts'; + val rec_fn_Ts = map (typ_subst_atomic (rec_result_Ts' ~~ rec_result_Ts)) rec_fn_Ts'; + + val rec_set_Ts = map (fn (T1, T2) => + rec_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts); + + val big_rec_name = big_name ^ "_rec_set"; + val rec_set_names' = + if length descr'' = 1 then [big_rec_name] else + map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int) + (1 upto (length descr'')); + val rec_set_names = map (Sign.full_bname thy10) rec_set_names'; + + val rec_fns = map (uncurry (mk_Free "f")) + (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts))); + val rec_sets' = map (fn c => list_comb (Free c, rec_fns)) + (rec_set_names' ~~ rec_set_Ts); + val rec_sets = map (fn c => list_comb (Const c, rec_fns)) + (rec_set_names ~~ rec_set_Ts); + + (* introduction rules for graph of recursion function *) + + val rec_preds = map (fn (a, T) => + Free (a, T --> HOLogic.boolT)) (pnames ~~ rec_result_Ts); + + fun mk_fresh3 rs [] = [] + | mk_fresh3 rs ((p as (ys, z)) :: yss) = List.concat (map (fn y as (_, T) => + List.mapPartial (fn ((_, (_, x)), r as (_, U)) => if z = x then NONE + else SOME (HOLogic.mk_Trueprop + (fresh_const T U $ Free y $ Free r))) rs) ys) @ + mk_fresh3 rs yss; + + (* FIXME: avoid collisions with other variable names? *) + val rec_ctxt = Free ("z", fsT'); + + fun make_rec_intr T p rec_set ((rec_intr_ts, rec_prems, rec_prems', + rec_eq_prems, l), ((cname, cargs), idxs)) = + let + val Ts = map (typ_of_dtyp descr'' sorts) cargs; + val frees = map (fn i => "x" ^ string_of_int i) (1 upto length Ts) ~~ Ts; + val frees' = partition_cargs idxs frees; + val binders = List.concat (map fst frees'); + val atomTs = distinct op = (maps (map snd o fst) frees'); + val recs = List.mapPartial + (fn ((_, DtRec i), p) => SOME (i, p) | _ => NONE) + (partition_cargs idxs cargs ~~ frees'); + val frees'' = map (fn i => "y" ^ string_of_int i) (1 upto length recs) ~~ + map (fn (i, _) => List.nth (rec_result_Ts, i)) recs; + val prems1 = map (fn ((i, (_, x)), y) => HOLogic.mk_Trueprop + (List.nth (rec_sets', i) $ Free x $ Free y)) (recs ~~ frees''); + val prems2 = + map (fn f => map (fn p as (_, T) => HOLogic.mk_Trueprop + (fresh_const T (fastype_of f) $ Free p $ f)) binders) rec_fns; + val prems3 = mk_fresh1 [] binders @ mk_fresh2 [] frees'; + val prems4 = map (fn ((i, _), y) => + HOLogic.mk_Trueprop (List.nth (rec_preds, i) $ Free y)) (recs ~~ frees''); + val prems5 = mk_fresh3 (recs ~~ frees'') frees'; + val prems6 = maps (fn aT => map (fn y as (_, T) => HOLogic.mk_Trueprop + (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $ + (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ Free y))) + frees'') atomTs; + val prems7 = map (fn x as (_, T) => HOLogic.mk_Trueprop + (fresh_const T fsT' $ Free x $ rec_ctxt)) binders; + val result = list_comb (List.nth (rec_fns, l), map Free (frees @ frees'')); + val result_freshs = map (fn p as (_, T) => + fresh_const T (fastype_of result) $ Free p $ result) binders; + val P = HOLogic.mk_Trueprop (p $ result) + in + (rec_intr_ts @ [Logic.list_implies (List.concat prems2 @ prems3 @ prems1, + HOLogic.mk_Trueprop (rec_set $ + list_comb (Const (cname, Ts ---> T), map Free frees) $ result))], + rec_prems @ [list_all_free (frees @ frees'', Logic.list_implies (prems4, P))], + rec_prems' @ map (fn fr => list_all_free (frees @ frees'', + Logic.list_implies (List.nth (prems2, l) @ prems3 @ prems5 @ prems7 @ prems6 @ [P], + HOLogic.mk_Trueprop fr))) result_freshs, + rec_eq_prems @ [List.concat prems2 @ prems3], + l + 1) + end; + + val (rec_intr_ts, rec_prems, rec_prems', rec_eq_prems, _) = + Library.foldl (fn (x, ((((d, d'), T), p), rec_set)) => + Library.foldl (make_rec_intr T p rec_set) (x, #3 (snd d) ~~ snd d')) + (([], [], [], [], 0), descr'' ~~ ndescr ~~ recTs ~~ rec_preds ~~ rec_sets'); + + val ({intrs = rec_intrs, elims = rec_elims, raw_induct = rec_induct, ...}, thy11) = + thy10 |> + Inductive.add_inductive_global (serial_string ()) + {quiet_mode = #quiet config, verbose = false, kind = Thm.internalK, + alt_name = Binding.name big_rec_name, coind = false, no_elim = false, no_ind = false, + skip_mono = true, fork_mono = false} + (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts)) + (map dest_Free rec_fns) + (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] ||> + PureThy.hide_fact true (Long_Name.append (Sign.full_bname thy10 big_rec_name) "induct"); + + (** equivariance **) + + val fresh_bij = PureThy.get_thms thy11 "fresh_bij"; + val perm_bij = PureThy.get_thms thy11 "perm_bij"; + + val (rec_equiv_thms, rec_equiv_thms') = ListPair.unzip (map (fn aT => + let + val permT = mk_permT aT; + val pi = Free ("pi", permT); + val rec_fns_pi = map (mk_perm [] pi o uncurry (mk_Free "f")) + (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts))); + val rec_sets_pi = map (fn c => list_comb (Const c, rec_fns_pi)) + (rec_set_names ~~ rec_set_Ts); + val ps = map (fn ((((T, U), R), R'), i) => + let + val x = Free ("x" ^ string_of_int i, T); + val y = Free ("y" ^ string_of_int i, U) + in + (R $ x $ y, R' $ mk_perm [] pi x $ mk_perm [] pi y) + end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ rec_sets_pi ~~ (1 upto length recTs)); + val ths = map (fn th => standard (th RS mp)) (split_conj_thm + (Goal.prove_global thy11 [] [] + (augment_sort thy1 pt_cp_sort + (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps)))) + (fn _ => rtac rec_induct 1 THEN REPEAT + (simp_tac (Simplifier.theory_context thy11 HOL_basic_ss + addsimps flat perm_simps' + addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN + (resolve_tac rec_intrs THEN_ALL_NEW + asm_simp_tac (HOL_ss addsimps (fresh_bij @ perm_bij))) 1)))) + val ths' = map (fn ((P, Q), th) => + Goal.prove_global thy11 [] [] + (augment_sort thy1 pt_cp_sort + (Logic.mk_implies (HOLogic.mk_Trueprop Q, HOLogic.mk_Trueprop P))) + (fn _ => dtac (Thm.instantiate ([], + [(cterm_of thy11 (Var (("pi", 0), permT)), + cterm_of thy11 (Const ("List.rev", permT --> permT) $ pi))]) th) 1 THEN + NominalPermeq.perm_simp_tac HOL_ss 1)) (ps ~~ ths) + in (ths, ths') end) dt_atomTs); + + (** finite support **) + + val rec_fin_supp_thms = map (fn aT => + let + val name = Long_Name.base_name (fst (dest_Type aT)); + val fs_name = PureThy.get_thm thy11 ("fs_" ^ name ^ "1"); + val aset = HOLogic.mk_setT aT; + val finite = Const ("Finite_Set.finite", aset --> HOLogic.boolT); + val fins = map (fn (f, T) => HOLogic.mk_Trueprop + (finite $ (Const ("Nominal.supp", T --> aset) $ f))) + (rec_fns ~~ rec_fn_Ts) + in + map (fn th => standard (th RS mp)) (split_conj_thm + (Goal.prove_global thy11 [] + (map (augment_sort thy11 fs_cp_sort) fins) + (augment_sort thy11 fs_cp_sort + (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj + (map (fn (((T, U), R), i) => + let + val x = Free ("x" ^ string_of_int i, T); + val y = Free ("y" ^ string_of_int i, U) + in + HOLogic.mk_imp (R $ x $ y, + finite $ (Const ("Nominal.supp", U --> aset) $ y)) + end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ + (1 upto length recTs)))))) + (fn {prems = fins, ...} => + (rtac rec_induct THEN_ALL_NEW cut_facts_tac fins) 1 THEN REPEAT + (NominalPermeq.finite_guess_tac (HOL_ss addsimps [fs_name]) 1)))) + end) dt_atomTs; + + (** freshness **) + + val finite_premss = map (fn aT => + map (fn (f, T) => HOLogic.mk_Trueprop + (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $ + (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ f))) + (rec_fns ~~ rec_fn_Ts)) dt_atomTs; + + val rec_fns' = map (augment_sort thy11 fs_cp_sort) rec_fns; + + val rec_fresh_thms = map (fn ((aT, eqvt_ths), finite_prems) => + let + val name = Long_Name.base_name (fst (dest_Type aT)); + val fs_name = PureThy.get_thm thy11 ("fs_" ^ name ^ "1"); + val a = Free ("a", aT); + val freshs = map (fn (f, fT) => HOLogic.mk_Trueprop + (fresh_const aT fT $ a $ f)) (rec_fns ~~ rec_fn_Ts) + in + map (fn (((T, U), R), eqvt_th) => + let + val x = Free ("x", augment_sort_typ thy11 fs_cp_sort T); + val y = Free ("y", U); + val y' = Free ("y'", U) + in + standard (Goal.prove (ProofContext.init thy11) [] + (map (augment_sort thy11 fs_cp_sort) + (finite_prems @ + [HOLogic.mk_Trueprop (R $ x $ y), + HOLogic.mk_Trueprop (HOLogic.mk_all ("y'", U, + HOLogic.mk_imp (R $ x $ y', HOLogic.mk_eq (y', y)))), + HOLogic.mk_Trueprop (fresh_const aT T $ a $ x)] @ + freshs)) + (HOLogic.mk_Trueprop (fresh_const aT U $ a $ y)) + (fn {prems, context} => + let + val (finite_prems, rec_prem :: unique_prem :: + fresh_prems) = chop (length finite_prems) prems; + val unique_prem' = unique_prem RS spec RS mp; + val unique = [unique_prem', unique_prem' RS sym] MRS trans; + val _ $ (_ $ (_ $ S $ _)) $ _ = prop_of supports_fresh; + val tuple = foldr1 HOLogic.mk_prod (x :: rec_fns') + in EVERY + [rtac (Drule.cterm_instantiate + [(cterm_of thy11 S, + cterm_of thy11 (Const ("Nominal.supp", + fastype_of tuple --> HOLogic.mk_setT aT) $ tuple))] + supports_fresh) 1, + simp_tac (HOL_basic_ss addsimps + [supports_def, symmetric fresh_def, fresh_prod]) 1, + REPEAT_DETERM (resolve_tac [allI, impI] 1), + REPEAT_DETERM (etac conjE 1), + rtac unique 1, + SUBPROOF (fn {prems = prems', params = [a, b], ...} => EVERY + [cut_facts_tac [rec_prem] 1, + rtac (Thm.instantiate ([], + [(cterm_of thy11 (Var (("pi", 0), mk_permT aT)), + cterm_of thy11 (perm_of_pair (term_of a, term_of b)))]) eqvt_th) 1, + asm_simp_tac (HOL_ss addsimps + (prems' @ perm_swap @ perm_fresh_fresh)) 1]) context 1, + rtac rec_prem 1, + simp_tac (HOL_ss addsimps (fs_name :: + supp_prod :: finite_Un :: finite_prems)) 1, + simp_tac (HOL_ss addsimps (symmetric fresh_def :: + fresh_prod :: fresh_prems)) 1] + end)) + end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ eqvt_ths) + end) (dt_atomTs ~~ rec_equiv_thms' ~~ finite_premss); + + (** uniqueness **) + + val fun_tuple = foldr1 HOLogic.mk_prod (rec_ctxt :: rec_fns); + val fun_tupleT = fastype_of fun_tuple; + val rec_unique_frees = + DatatypeProp.indexify_names (replicate (length recTs) "x") ~~ recTs; + val rec_unique_frees'' = map (fn (s, T) => (s ^ "'", T)) rec_unique_frees; + val rec_unique_frees' = + DatatypeProp.indexify_names (replicate (length recTs) "y") ~~ rec_result_Ts; + val rec_unique_concls = map (fn ((x, U), R) => + Const ("Ex1", (U --> HOLogic.boolT) --> HOLogic.boolT) $ + Abs ("y", U, R $ Free x $ Bound 0)) + (rec_unique_frees ~~ rec_result_Ts ~~ rec_sets); + + val induct_aux_rec = Drule.cterm_instantiate + (map (pairself (cterm_of thy11) o apsnd (augment_sort thy11 fs_cp_sort)) + (map (fn (aT, f) => (Logic.varify f, Abs ("z", HOLogic.unitT, + Const ("Nominal.supp", fun_tupleT --> HOLogic.mk_setT aT) $ fun_tuple))) + fresh_fs @ + map (fn (((P, T), (x, U)), Q) => + (Var ((P, 0), Logic.varifyT (fsT' --> T --> HOLogic.boolT)), + Abs ("z", HOLogic.unitT, absfree (x, U, Q)))) + (pnames ~~ recTs ~~ rec_unique_frees ~~ rec_unique_concls) @ + map (fn (s, T) => (Var ((s, 0), Logic.varifyT T), Free (s, T))) + rec_unique_frees)) induct_aux; + + fun obtain_fresh_name vs ths rec_fin_supp T (freshs1, freshs2, ctxt) = + let + val p = foldr1 HOLogic.mk_prod (vs @ freshs1); + val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop + (HOLogic.exists_const T $ Abs ("x", T, + fresh_const T (fastype_of p) $ Bound 0 $ p))) + (fn _ => EVERY + [cut_facts_tac ths 1, + REPEAT_DETERM (dresolve_tac (the (AList.lookup op = rec_fin_supp T)) 1), + resolve_tac exists_fresh' 1, + asm_simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms)) 1]); + val (([cx], ths), ctxt') = Obtain.result + (fn _ => EVERY + [etac exE 1, + full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1, + REPEAT (etac conjE 1)]) + [ex] ctxt + in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end; + + val finite_ctxt_prems = map (fn aT => + HOLogic.mk_Trueprop + (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $ + (Const ("Nominal.supp", fsT' --> HOLogic.mk_setT aT) $ rec_ctxt))) dt_atomTs; + + val rec_unique_thms = split_conj_thm (Goal.prove + (ProofContext.init thy11) (map fst rec_unique_frees) + (map (augment_sort thy11 fs_cp_sort) + (List.concat finite_premss @ finite_ctxt_prems @ rec_prems @ rec_prems')) + (augment_sort thy11 fs_cp_sort + (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj rec_unique_concls))) + (fn {prems, context} => + let + val k = length rec_fns; + val (finite_thss, ths1) = fold_map (fn T => fn xs => + apfst (pair T) (chop k xs)) dt_atomTs prems; + val (finite_ctxt_ths, ths2) = chop (length dt_atomTs) ths1; + val (P_ind_ths, fcbs) = chop k ths2; + val P_ths = map (fn th => th RS mp) (split_conj_thm + (Goal.prove context + (map fst (rec_unique_frees'' @ rec_unique_frees')) [] + (augment_sort thy11 fs_cp_sort + (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj + (map (fn (((x, y), S), P) => HOLogic.mk_imp + (S $ Free x $ Free y, P $ (Free y))) + (rec_unique_frees'' ~~ rec_unique_frees' ~~ + rec_sets ~~ rec_preds))))) + (fn _ => + rtac rec_induct 1 THEN + REPEAT ((resolve_tac P_ind_ths THEN_ALL_NEW assume_tac) 1)))); + val rec_fin_supp_thms' = map + (fn (ths, (T, fin_ths)) => (T, map (curry op MRS fin_ths) ths)) + (rec_fin_supp_thms ~~ finite_thss); + in EVERY + ([rtac induct_aux_rec 1] @ + maps (fn ((_, finite_ths), finite_th) => + [cut_facts_tac (finite_th :: finite_ths) 1, + asm_simp_tac (HOL_ss addsimps [supp_prod, finite_Un]) 1]) + (finite_thss ~~ finite_ctxt_ths) @ + maps (fn ((_, idxss), elim) => maps (fn idxs => + [full_simp_tac (HOL_ss addsimps [symmetric fresh_def, supp_prod, Un_iff]) 1, + REPEAT_DETERM (eresolve_tac [conjE, ex1E] 1), + rtac ex1I 1, + (resolve_tac rec_intrs THEN_ALL_NEW atac) 1, + rotate_tac ~1 1, + ((DETERM o etac elim) THEN_ALL_NEW full_simp_tac + (HOL_ss addsimps List.concat distinct_thms)) 1] @ + (if null idxs then [] else [hyp_subst_tac 1, + SUBPROOF (fn {asms, concl, prems = prems', params, context = context', ...} => + let + val SOME prem = find_first (can (HOLogic.dest_eq o + HOLogic.dest_Trueprop o prop_of)) prems'; + val _ $ (_ $ lhs $ rhs) = prop_of prem; + val _ $ (_ $ lhs' $ rhs') = term_of concl; + val rT = fastype_of lhs'; + val (c, cargsl) = strip_comb lhs; + val cargsl' = partition_cargs idxs cargsl; + val boundsl = List.concat (map fst cargsl'); + val (_, cargsr) = strip_comb rhs; + val cargsr' = partition_cargs idxs cargsr; + val boundsr = List.concat (map fst cargsr'); + val (params1, _ :: params2) = + chop (length params div 2) (map term_of params); + val params' = params1 @ params2; + val rec_prems = filter (fn th => case prop_of th of + _ $ p => (case head_of p of + Const (s, _) => s mem rec_set_names + | _ => false) + | _ => false) prems'; + val fresh_prems = filter (fn th => case prop_of th of + _ $ (Const ("Nominal.fresh", _) $ _ $ _) => true + | _ $ (Const ("Not", _) $ _) => true + | _ => false) prems'; + val Ts = map fastype_of boundsl; + + val _ = warning "step 1: obtaining fresh names"; + val (freshs1, freshs2, context'') = fold + (obtain_fresh_name (rec_ctxt :: rec_fns' @ params') + (List.concat (map snd finite_thss) @ + finite_ctxt_ths @ rec_prems) + rec_fin_supp_thms') + Ts ([], [], context'); + val pi1 = map perm_of_pair (boundsl ~~ freshs1); + val rpi1 = rev pi1; + val pi2 = map perm_of_pair (boundsr ~~ freshs1); + val rpi2 = rev pi2; + + val fresh_prems' = mk_not_sym fresh_prems; + val freshs2' = mk_not_sym freshs2; + + (** as, bs, cs # K as ts, K bs us **) + val _ = warning "step 2: as, bs, cs # K as ts, K bs us"; + val prove_fresh_ss = HOL_ss addsimps + (finite_Diff :: List.concat fresh_thms @ + fs_atoms @ abs_fresh @ abs_supp @ fresh_atm); + (* FIXME: avoid asm_full_simp_tac ? *) + fun prove_fresh ths y x = Goal.prove context'' [] [] + (HOLogic.mk_Trueprop (fresh_const + (fastype_of x) (fastype_of y) $ x $ y)) + (fn _ => cut_facts_tac ths 1 THEN asm_full_simp_tac prove_fresh_ss 1); + val constr_fresh_thms = + map (prove_fresh fresh_prems lhs) boundsl @ + map (prove_fresh fresh_prems rhs) boundsr @ + map (prove_fresh freshs2 lhs) freshs1 @ + map (prove_fresh freshs2 rhs) freshs1; + + (** pi1 o (K as ts) = pi2 o (K bs us) **) + val _ = warning "step 3: pi1 o (K as ts) = pi2 o (K bs us)"; + val pi1_pi2_eq = Goal.prove context'' [] [] + (HOLogic.mk_Trueprop (HOLogic.mk_eq + (fold_rev (mk_perm []) pi1 lhs, fold_rev (mk_perm []) pi2 rhs))) + (fn _ => EVERY + [cut_facts_tac constr_fresh_thms 1, + asm_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh) 1, + rtac prem 1]); + + (** pi1 o ts = pi2 o us **) + val _ = warning "step 4: pi1 o ts = pi2 o us"; + val pi1_pi2_eqs = map (fn (t, u) => + Goal.prove context'' [] [] + (HOLogic.mk_Trueprop (HOLogic.mk_eq + (fold_rev (mk_perm []) pi1 t, fold_rev (mk_perm []) pi2 u))) + (fn _ => EVERY + [cut_facts_tac [pi1_pi2_eq] 1, + asm_full_simp_tac (HOL_ss addsimps + (calc_atm @ List.concat perm_simps' @ + fresh_prems' @ freshs2' @ abs_perm @ + alpha @ List.concat inject_thms)) 1])) + (map snd cargsl' ~~ map snd cargsr'); + + (** pi1^-1 o pi2 o us = ts **) + val _ = warning "step 5: pi1^-1 o pi2 o us = ts"; + val rpi1_pi2_eqs = map (fn ((t, u), eq) => + Goal.prove context'' [] [] + (HOLogic.mk_Trueprop (HOLogic.mk_eq + (fold_rev (mk_perm []) (rpi1 @ pi2) u, t))) + (fn _ => simp_tac (HOL_ss addsimps + ((eq RS sym) :: perm_swap)) 1)) + (map snd cargsl' ~~ map snd cargsr' ~~ pi1_pi2_eqs); + + val (rec_prems1, rec_prems2) = + chop (length rec_prems div 2) rec_prems; + + (** (ts, pi1^-1 o pi2 o vs) in rec_set **) + val _ = warning "step 6: (ts, pi1^-1 o pi2 o vs) in rec_set"; + val rec_prems' = map (fn th => + let + val _ $ (S $ x $ y) = prop_of th; + val Const (s, _) = head_of S; + val k = find_index (equal s) rec_set_names; + val pi = rpi1 @ pi2; + fun mk_pi z = fold_rev (mk_perm []) pi z; + fun eqvt_tac p = + let + val U as Type (_, [Type (_, [T, _])]) = fastype_of p; + val l = find_index (equal T) dt_atomTs; + val th = List.nth (List.nth (rec_equiv_thms', l), k); + val th' = Thm.instantiate ([], + [(cterm_of thy11 (Var (("pi", 0), U)), + cterm_of thy11 p)]) th; + in rtac th' 1 end; + val th' = Goal.prove context'' [] [] + (HOLogic.mk_Trueprop (S $ mk_pi x $ mk_pi y)) + (fn _ => EVERY + (map eqvt_tac pi @ + [simp_tac (HOL_ss addsimps (fresh_prems' @ freshs2' @ + perm_swap @ perm_fresh_fresh)) 1, + rtac th 1])) + in + Simplifier.simplify + (HOL_basic_ss addsimps rpi1_pi2_eqs) th' + end) rec_prems2; + + val ihs = filter (fn th => case prop_of th of + _ $ (Const ("All", _) $ _) => true | _ => false) prems'; + + (** pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs **) + val _ = warning "step 7: pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs"; + val rec_eqns = map (fn (th, ih) => + let + val th' = th RS (ih RS spec RS mp) RS sym; + val _ $ (_ $ lhs $ rhs) = prop_of th'; + fun strip_perm (_ $ _ $ t) = strip_perm t + | strip_perm t = t; + in + Goal.prove context'' [] [] + (HOLogic.mk_Trueprop (HOLogic.mk_eq + (fold_rev (mk_perm []) pi1 lhs, + fold_rev (mk_perm []) pi2 (strip_perm rhs)))) + (fn _ => simp_tac (HOL_basic_ss addsimps + (th' :: perm_swap)) 1) + end) (rec_prems' ~~ ihs); + + (** as # rs **) + val _ = warning "step 8: as # rs"; + val rec_freshs = List.concat + (map (fn (rec_prem, ih) => + let + val _ $ (S $ x $ (y as Free (_, T))) = + prop_of rec_prem; + val k = find_index (equal S) rec_sets; + val atoms = List.concat (List.mapPartial (fn (bs, z) => + if z = x then NONE else SOME bs) cargsl') + in + map (fn a as Free (_, aT) => + let val l = find_index (equal aT) dt_atomTs; + in + Goal.prove context'' [] [] + (HOLogic.mk_Trueprop (fresh_const aT T $ a $ y)) + (fn _ => EVERY + (rtac (List.nth (List.nth (rec_fresh_thms, l), k)) 1 :: + map (fn th => rtac th 1) + (snd (List.nth (finite_thss, l))) @ + [rtac rec_prem 1, rtac ih 1, + REPEAT_DETERM (resolve_tac fresh_prems 1)])) + end) atoms + end) (rec_prems1 ~~ ihs)); + + (** as # fK as ts rs , bs # fK bs us vs **) + val _ = warning "step 9: as # fK as ts rs , bs # fK bs us vs"; + fun prove_fresh_result (a as Free (_, aT)) = + Goal.prove context'' [] [] + (HOLogic.mk_Trueprop (fresh_const aT rT $ a $ rhs')) + (fn _ => EVERY + [resolve_tac fcbs 1, + REPEAT_DETERM (resolve_tac + (fresh_prems @ rec_freshs) 1), + REPEAT_DETERM (resolve_tac (maps snd rec_fin_supp_thms') 1 + THEN resolve_tac rec_prems 1), + resolve_tac P_ind_ths 1, + REPEAT_DETERM (resolve_tac (P_ths @ rec_prems) 1)]); + + val fresh_results'' = map prove_fresh_result boundsl; + + fun prove_fresh_result'' ((a as Free (_, aT), b), th) = + let val th' = Goal.prove context'' [] [] + (HOLogic.mk_Trueprop (fresh_const aT rT $ + fold_rev (mk_perm []) (rpi2 @ pi1) a $ + fold_rev (mk_perm []) (rpi2 @ pi1) rhs')) + (fn _ => simp_tac (HOL_ss addsimps fresh_bij) 1 THEN + rtac th 1) + in + Goal.prove context'' [] [] + (HOLogic.mk_Trueprop (fresh_const aT rT $ b $ lhs')) + (fn _ => EVERY + [cut_facts_tac [th'] 1, + full_simp_tac (Simplifier.theory_context thy11 HOL_ss + addsimps rec_eqns @ pi1_pi2_eqs @ perm_swap + addsimprocs [NominalPermeq.perm_simproc_app]) 1, + full_simp_tac (HOL_ss addsimps (calc_atm @ + fresh_prems' @ freshs2' @ perm_fresh_fresh)) 1]) + end; + + val fresh_results = fresh_results'' @ map prove_fresh_result'' + (boundsl ~~ boundsr ~~ fresh_results''); + + (** cs # fK as ts rs , cs # fK bs us vs **) + val _ = warning "step 10: cs # fK as ts rs , cs # fK bs us vs"; + fun prove_fresh_result' recs t (a as Free (_, aT)) = + Goal.prove context'' [] [] + (HOLogic.mk_Trueprop (fresh_const aT rT $ a $ t)) + (fn _ => EVERY + [cut_facts_tac recs 1, + REPEAT_DETERM (dresolve_tac + (the (AList.lookup op = rec_fin_supp_thms' aT)) 1), + NominalPermeq.fresh_guess_tac + (HOL_ss addsimps (freshs2 @ + fs_atoms @ fresh_atm @ + List.concat (map snd finite_thss))) 1]); + + val fresh_results' = + map (prove_fresh_result' rec_prems1 rhs') freshs1 @ + map (prove_fresh_result' rec_prems2 lhs') freshs1; + + (** pi1 o (fK as ts rs) = pi2 o (fK bs us vs) **) + val _ = warning "step 11: pi1 o (fK as ts rs) = pi2 o (fK bs us vs)"; + val pi1_pi2_result = Goal.prove context'' [] [] + (HOLogic.mk_Trueprop (HOLogic.mk_eq + (fold_rev (mk_perm []) pi1 rhs', fold_rev (mk_perm []) pi2 lhs'))) + (fn _ => simp_tac (Simplifier.context context'' HOL_ss + addsimps pi1_pi2_eqs @ rec_eqns + addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN + TRY (simp_tac (HOL_ss addsimps + (fresh_prems' @ freshs2' @ calc_atm @ perm_fresh_fresh)) 1)); + + val _ = warning "final result"; + val final = Goal.prove context'' [] [] (term_of concl) + (fn _ => cut_facts_tac [pi1_pi2_result RS sym] 1 THEN + full_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh @ + fresh_results @ fresh_results') 1); + val final' = ProofContext.export context'' context' [final]; + val _ = warning "finished!" + in + resolve_tac final' 1 + end) context 1])) idxss) (ndescr ~~ rec_elims)) + end)); + + val rec_total_thms = map (fn r => r RS theI') rec_unique_thms; + + (* define primrec combinators *) + + val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec"; + val reccomb_names = map (Sign.full_bname thy11) + (if length descr'' = 1 then [big_reccomb_name] else + (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int) + (1 upto (length descr'')))); + val reccombs = map (fn ((name, T), T') => list_comb + (Const (name, rec_fn_Ts @ [T] ---> T'), rec_fns)) + (reccomb_names ~~ recTs ~~ rec_result_Ts); + + val (reccomb_defs, thy12) = + thy11 + |> Sign.add_consts_i (map (fn ((name, T), T') => + (Binding.name (Long_Name.base_name name), rec_fn_Ts @ [T] ---> T', NoSyn)) + (reccomb_names ~~ recTs ~~ rec_result_Ts)) + |> (PureThy.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') => + (Binding.name (Long_Name.base_name name ^ "_def"), Logic.mk_equals (comb, absfree ("x", T, + Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T', + set $ Free ("x", T) $ Free ("y", T')))))) + (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts)); + + (* prove characteristic equations for primrec combinators *) + + val rec_thms = map (fn (prems, concl) => + let + val _ $ (_ $ (_ $ x) $ _) = concl; + val (_, cargs) = strip_comb x; + val ps = map (fn (x as Free (_, T), i) => + (Free ("x" ^ string_of_int i, T), x)) (cargs ~~ (1 upto length cargs)); + val concl' = subst_atomic_types (rec_result_Ts' ~~ rec_result_Ts) concl; + val prems' = List.concat finite_premss @ finite_ctxt_prems @ + rec_prems @ rec_prems' @ map (subst_atomic ps) prems; + fun solve rules prems = resolve_tac rules THEN_ALL_NEW + (resolve_tac prems THEN_ALL_NEW atac) + in + Goal.prove_global thy12 [] + (map (augment_sort thy12 fs_cp_sort) prems') + (augment_sort thy12 fs_cp_sort concl') + (fn {prems, ...} => EVERY + [rewrite_goals_tac reccomb_defs, + rtac the1_equality 1, + solve rec_unique_thms prems 1, + resolve_tac rec_intrs 1, + REPEAT (solve (prems @ rec_total_thms) prems 1)]) + end) (rec_eq_prems ~~ + DatatypeProp.make_primrecs new_type_names descr' sorts thy12); + + val dt_infos = map (make_dt_info pdescr sorts induct reccomb_names rec_thms) + ((0 upto length descr1 - 1) ~~ descr1 ~~ distinct_thms ~~ inject_thms); + + (* FIXME: theorems are stored in database for testing only *) + val (_, thy13) = thy12 |> + PureThy.add_thmss + [((Binding.name "rec_equiv", List.concat rec_equiv_thms), []), + ((Binding.name "rec_equiv'", List.concat rec_equiv_thms'), []), + ((Binding.name "rec_fin_supp", List.concat rec_fin_supp_thms), []), + ((Binding.name "rec_fresh", List.concat rec_fresh_thms), []), + ((Binding.name "rec_unique", map standard rec_unique_thms), []), + ((Binding.name "recs", rec_thms), [])] ||> + Sign.parent_path ||> + map_nominal_datatypes (fold Symtab.update dt_infos); + + in + thy13 + end; + +val add_nominal_datatype = gen_add_nominal_datatype Datatype.read_typ; + + +(* FIXME: The following stuff should be exported by Datatype *) + +local structure P = OuterParse and K = OuterKeyword in + +val datatype_decl = + Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix -- + (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix)); + +fun mk_datatype args = + let + val names = map (fn ((((NONE, _), t), _), _) => t | ((((SOME t, _), _), _), _) => t) args; + val specs = map (fn ((((_, vs), t), mx), cons) => + (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args; + in add_nominal_datatype Datatype.default_config names specs end; + +val _ = + OuterSyntax.command "nominal_datatype" "define inductive datatypes" K.thy_decl + (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype)); + +end; + +end diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Nominal/nominal_inductive.ML --- a/src/HOL/Nominal/nominal_inductive.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Nominal/nominal_inductive.ML Sat Jul 04 15:19:47 2009 +0200 @@ -53,7 +53,7 @@ fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of (Const (s, T), ts) => (case strip_type T of (Ts, Type (tname, _)) => - (case Nominal.get_nominal_datatype thy tname of + (case NominalDatatype.get_nominal_datatype thy tname of NONE => fold (add_binders thy i) ts bs | SOME {descr, index, ...} => (case AList.lookup op = (#3 (the (AList.lookup op = descr index))) s of @@ -230,7 +230,7 @@ else NONE) xs @ mk_distinct xs; fun mk_fresh (x, T) = HOLogic.mk_Trueprop - (Nominal.fresh_const T fsT $ x $ Bound 0); + (NominalDatatype.fresh_const T fsT $ x $ Bound 0); val (prems', prems'') = split_list (map (fn (params, bvars, prems, (p, ts)) => let @@ -254,7 +254,7 @@ val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, HOLogic.list_all (ind_vars, lift_pred p - (map (fold_rev (Nominal.mk_perm ind_Ts) + (map (fold_rev (NominalDatatype.mk_perm ind_Ts) (map Bound (length atomTs downto 1))) ts)))) concls)); val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj @@ -268,7 +268,7 @@ else map_term (split_conj (K o I) names) prem prem) prems, q)))) (mk_distinct bvars @ maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop - (Nominal.fresh_const U T $ u $ t)) bvars) + (NominalDatatype.fresh_const U T $ u $ t)) bvars) (ts ~~ binder_types (fastype_of p)))) prems; val perm_pi_simp = PureThy.get_thms thy "perm_pi_simp"; @@ -296,7 +296,7 @@ val p = foldr1 HOLogic.mk_prod (map protect ts @ freshs1); val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.exists_const T $ Abs ("x", T, - Nominal.fresh_const T (fastype_of p) $ + NominalDatatype.fresh_const T (fastype_of p) $ Bound 0 $ p))) (fn _ => EVERY [resolve_tac exists_fresh' 1, @@ -325,13 +325,13 @@ (fn (Bound i, T) => (nth params' (length params' - i), T) | (t, T) => (t, T)) bvars; val pi_bvars = map (fn (t, _) => - fold_rev (Nominal.mk_perm []) pis t) bvars'; + fold_rev (NominalDatatype.mk_perm []) pis t) bvars'; val (P, ts) = strip_comb (HOLogic.dest_Trueprop (term_of concl)); val (freshs1, freshs2, ctxt'') = fold (obtain_fresh_name (ts @ pi_bvars)) (map snd bvars') ([], [], ctxt'); - val freshs2' = Nominal.mk_not_sym freshs2; - val pis' = map Nominal.perm_of_pair (pi_bvars ~~ freshs1); + val freshs2' = NominalDatatype.mk_not_sym freshs2; + val pis' = map NominalDatatype.perm_of_pair (pi_bvars ~~ freshs1); fun concat_perm pi1 pi2 = let val T = fastype_of pi1 in if T = fastype_of pi2 then @@ -343,11 +343,11 @@ (Vartab.empty, Vartab.empty); val ihyp' = Thm.instantiate ([], map (pairself (cterm_of thy)) (map (Envir.subst_vars env) vs ~~ - map (fold_rev (Nominal.mk_perm []) + map (fold_rev (NominalDatatype.mk_perm []) (rev pis' @ pis)) params' @ [z])) ihyp; fun mk_pi th = Simplifier.simplify (HOL_basic_ss addsimps [@{thm id_apply}] - addsimprocs [Nominal.perm_simproc]) + addsimprocs [NominalDatatype.perm_simproc]) (Simplifier.simplify eqvt_ss (fold_rev (mk_perm_bool o cterm_of thy) (rev pis' @ pis) th)); @@ -369,13 +369,13 @@ | _ $ (_ $ (_ $ lhs $ rhs)) => (curry (HOLogic.mk_not o HOLogic.mk_eq), lhs, rhs)); val th'' = Goal.prove ctxt'' [] [] (HOLogic.mk_Trueprop - (bop (fold_rev (Nominal.mk_perm []) pis lhs) - (fold_rev (Nominal.mk_perm []) pis rhs))) + (bop (fold_rev (NominalDatatype.mk_perm []) pis lhs) + (fold_rev (NominalDatatype.mk_perm []) pis rhs))) (fn _ => simp_tac (HOL_basic_ss addsimps (fresh_bij @ perm_bij)) 1 THEN rtac th' 1) in Simplifier.simplify (eqvt_ss addsimps fresh_atm) th'' end) vc_compat_ths; - val vc_compat_ths'' = Nominal.mk_not_sym vc_compat_ths'; + val vc_compat_ths'' = NominalDatatype.mk_not_sym vc_compat_ths'; (** Since swap_simps simplifies (pi :: 'a prm) o (x :: 'b) to x **) (** we have to pre-simplify the rewrite rules **) val swap_simps_ss = HOL_ss addsimps swap_simps @ @@ -383,14 +383,14 @@ (vc_compat_ths'' @ freshs2'); val th = Goal.prove ctxt'' [] [] (HOLogic.mk_Trueprop (list_comb (P $ hd ts, - map (fold (Nominal.mk_perm []) pis') (tl ts)))) + map (fold (NominalDatatype.mk_perm []) pis') (tl ts)))) (fn _ => EVERY ([simp_tac eqvt_ss 1, rtac ihyp' 1, REPEAT_DETERM_N (nprems_of ihyp - length gprems) (simp_tac swap_simps_ss 1), REPEAT_DETERM_N (length gprems) (simp_tac (HOL_basic_ss addsimps [inductive_forall_def'] - addsimprocs [Nominal.perm_simproc]) 1 THEN + addsimprocs [NominalDatatype.perm_simproc]) 1 THEN resolve_tac gprems2 1)])); val final = Goal.prove ctxt'' [] [] (term_of concl) (fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (HOL_ss @@ -448,7 +448,7 @@ (Logic.list_implies (mk_distinct qs @ maps (fn (t, T) => map (fn u => HOLogic.mk_Trueprop - (Nominal.fresh_const T (fastype_of u) $ t $ u)) + (NominalDatatype.fresh_const T (fastype_of u) $ t $ u)) args) qs, HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.dest_Trueprop prems))), @@ -499,13 +499,13 @@ chop (length vc_compat_ths - length args * length qs) (map (first_order_mrs hyps2) vc_compat_ths); val vc_compat_ths' = - Nominal.mk_not_sym vc_compat_ths1 @ + NominalDatatype.mk_not_sym vc_compat_ths1 @ flat (fst (fold_map inst_fresh hyps1 vc_compat_ths2)); val (freshs1, freshs2, ctxt3) = fold (obtain_fresh_name (args @ map fst qs @ params')) (map snd qs) ([], [], ctxt2); - val freshs2' = Nominal.mk_not_sym freshs2; - val pis = map (Nominal.perm_of_pair) + val freshs2' = NominalDatatype.mk_not_sym freshs2; + val pis = map (NominalDatatype.perm_of_pair) ((freshs1 ~~ map fst qs) @ (params' ~~ freshs1)); val mk_pis = fold_rev mk_perm_bool (map (cterm_of thy) pis); val obj = cterm_of thy (foldr1 HOLogic.mk_conj (map (map_aterms @@ -513,7 +513,7 @@ if x mem args then x else (case AList.lookup op = tab x of SOME y => y - | NONE => fold_rev (Nominal.mk_perm []) pis x) + | NONE => fold_rev (NominalDatatype.mk_perm []) pis x) | x => x) o HOLogic.dest_Trueprop o prop_of) case_hyps)); val inst = Thm.first_order_match (Thm.dest_arg (Drule.strip_imp_concl (hd (cprems_of case_hyp))), obj); @@ -522,7 +522,7 @@ rtac (Thm.instantiate inst case_hyp) 1 THEN SUBPROOF (fn {prems = fresh_hyps, ...} => let - val fresh_hyps' = Nominal.mk_not_sym fresh_hyps; + val fresh_hyps' = NominalDatatype.mk_not_sym fresh_hyps; val case_ss = cases_eqvt_ss addsimps freshs2' @ simp_fresh_atm (vc_compat_ths' @ fresh_hyps'); val fresh_fresh_ss = case_ss addsimps perm_fresh_fresh; @@ -635,7 +635,7 @@ val prems'' = map (fn th => Simplifier.simplify eqvt_ss (mk_perm_bool (cterm_of thy pi) th)) prems'; val intr' = Drule.cterm_instantiate (map (cterm_of thy) vs ~~ - map (cterm_of thy o Nominal.mk_perm [] pi o term_of) params) + map (cterm_of thy o NominalDatatype.mk_perm [] pi o term_of) params) intr in (rtac intr' THEN_ALL_NEW (TRY o resolve_tac prems'')) 1 end) ctxt' 1 st @@ -655,7 +655,7 @@ val (ts1, ts2) = chop k ts in HOLogic.mk_imp (p, list_comb (h, ts1 @ - map (Nominal.mk_perm [] pi') ts2)) + map (NominalDatatype.mk_perm [] pi') ts2)) end) ps))) (fn {context, ...} => EVERY (rtac raw_induct 1 :: map (fn intr_vs => full_simp_tac eqvt_ss 1 THEN diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Nominal/nominal_inductive2.ML --- a/src/HOL/Nominal/nominal_inductive2.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Nominal/nominal_inductive2.ML Sat Jul 04 15:19:47 2009 +0200 @@ -56,7 +56,7 @@ fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of (Const (s, T), ts) => (case strip_type T of (Ts, Type (tname, _)) => - (case Nominal.get_nominal_datatype thy tname of + (case NominalDatatype.get_nominal_datatype thy tname of NONE => fold (add_binders thy i) ts bs | SOME {descr, index, ...} => (case AList.lookup op = (#3 (the (AList.lookup op = descr index))) s of @@ -249,7 +249,7 @@ | lift_prem t = t; fun mk_fresh (x, T) = HOLogic.mk_Trueprop - (Nominal.fresh_star_const T fsT $ x $ Bound 0); + (NominalDatatype.fresh_star_const T fsT $ x $ Bound 0); val (prems', prems'') = split_list (map (fn (params, sets, prems, (p, ts)) => let @@ -270,7 +270,7 @@ val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, HOLogic.list_all (ind_vars, lift_pred p - (map (fold_rev (Nominal.mk_perm ind_Ts) + (map (fold_rev (NominalDatatype.mk_perm ind_Ts) (map Bound (length atomTs downto 1))) ts)))) concls)); val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj @@ -283,7 +283,7 @@ if null (preds_of ps prem) then SOME prem else map_term (split_conj (K o I) names) prem prem) prems, q)))) (maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop - (Nominal.fresh_star_const U T $ u $ t)) sets) + (NominalDatatype.fresh_star_const U T $ u $ t)) sets) (ts ~~ binder_types (fastype_of p)) @ map (fn (u, U) => HOLogic.mk_Trueprop (Const (@{const_name finite}, HOLogic.mk_setT U --> HOLogic.boolT) $ u)) sets) |> @@ -339,7 +339,7 @@ val th2' = Goal.prove ctxt [] [] (list_all (map (pair "pi") pTs, HOLogic.mk_Trueprop - (f $ fold_rev (Nominal.mk_perm (rev pTs)) + (f $ fold_rev (NominalDatatype.mk_perm (rev pTs)) (pis1 @ pi :: pis2) l $ r))) (fn _ => cut_facts_tac [th2] 1 THEN full_simp_tac (HOL_basic_ss addsimps perm_set_forget) 1) |> @@ -364,7 +364,7 @@ val params' = map term_of cparams' val sets' = map (apfst (curry subst_bounds (rev params'))) sets; val pi_sets = map (fn (t, _) => - fold_rev (Nominal.mk_perm []) pis t) sets'; + fold_rev (NominalDatatype.mk_perm []) pis t) sets'; val (P, ts) = strip_comb (HOLogic.dest_Trueprop (term_of concl)); val gprems1 = List.mapPartial (fn (th, t) => if null (preds_of ps t) then SOME th @@ -380,7 +380,7 @@ in Goal.prove ctxt' [] [] (HOLogic.mk_Trueprop (list_comb (h, - map (fold_rev (Nominal.mk_perm []) pis) ts))) + map (fold_rev (NominalDatatype.mk_perm []) pis) ts))) (fn _ => simp_tac (HOL_basic_ss addsimps (fresh_star_bij @ finite_ineq)) 1 THEN rtac th' 1) end) vc_compat_ths vc_compat_vs; @@ -400,11 +400,11 @@ end; val pis'' = fold_rev (concat_perm #> map) pis' pis; val ihyp' = inst_params thy vs_ihypt ihyp - (map (fold_rev (Nominal.mk_perm []) + (map (fold_rev (NominalDatatype.mk_perm []) (pis' @ pis) #> cterm_of thy) params' @ [cterm_of thy z]); fun mk_pi th = Simplifier.simplify (HOL_basic_ss addsimps [@{thm id_apply}] - addsimprocs [Nominal.perm_simproc]) + addsimprocs [NominalDatatype.perm_simproc]) (Simplifier.simplify eqvt_ss (fold_rev (mk_perm_bool o cterm_of thy) (pis' @ pis) th)); @@ -419,13 +419,13 @@ (fresh_ths2 ~~ sets); val th = Goal.prove ctxt'' [] [] (HOLogic.mk_Trueprop (list_comb (P $ hd ts, - map (fold_rev (Nominal.mk_perm []) pis') (tl ts)))) + map (fold_rev (NominalDatatype.mk_perm []) pis') (tl ts)))) (fn _ => EVERY ([simp_tac eqvt_ss 1, rtac ihyp' 1] @ map (fn th => rtac th 1) fresh_ths3 @ [REPEAT_DETERM_N (length gprems) (simp_tac (HOL_basic_ss addsimps [inductive_forall_def'] - addsimprocs [Nominal.perm_simproc]) 1 THEN + addsimprocs [NominalDatatype.perm_simproc]) 1 THEN resolve_tac gprems2 1)])); val final = Goal.prove ctxt'' [] [] (term_of concl) (fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (HOL_ss diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Nominal/nominal_primrec.ML --- a/src/HOL/Nominal/nominal_primrec.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Nominal/nominal_primrec.ML Sat Jul 04 15:19:47 2009 +0200 @@ -223,7 +223,7 @@ (* find datatypes which contain all datatypes in tnames' *) -fun find_dts (dt_info : Nominal.nominal_datatype_info Symtab.table) _ [] = [] +fun find_dts (dt_info : NominalDatatype.nominal_datatype_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 nominal datatype") @@ -247,7 +247,7 @@ val eqns' = map unquantify spec' val eqns = fold_rev (process_eqn lthy (fn v => Variable.is_fixed lthy v orelse exists (fn ((w, _), _) => v = Binding.name_of w) fixes)) spec' []; - val dt_info = Nominal.get_nominal_datatypes (ProofContext.theory_of lthy); + val dt_info = NominalDatatype.get_nominal_datatypes (ProofContext.theory_of lthy); val lsrs :: lsrss = maps (fn (_, (_, _, eqns)) => map (fn (_, (ls, _, rs, _, _)) => ls @ rs) eqns) eqns val _ = @@ -373,7 +373,7 @@ lthy'' |> LocalTheory.note Thm.generatedK ((qualify (Binding.name "simps"), map (Attrib.internal o K) - [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add]), + [Simplifier.simp_add, Nitpick_Const_Simps.add]), maps snd simps') |> snd end) diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/OrderedGroup.thy --- a/src/HOL/OrderedGroup.thy Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/OrderedGroup.thy Sat Jul 04 15:19:47 2009 +0200 @@ -22,13 +22,15 @@ \end{itemize} *} -ML{* -structure AlgebraSimps = - NamedThmsFun(val name = "algebra_simps" - val description = "algebra simplification rules"); +ML {* +structure Algebra_Simps = Named_Thms +( + val name = "algebra_simps" + val description = "algebra simplification rules" +) *} -setup AlgebraSimps.setup +setup Algebra_Simps.setup text{* The rewrites accumulated in @{text algebra_simps} deal with the classical algebraic structures of groups, rings and family. They simplify diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Predicate.thy --- a/src/HOL/Predicate.thy Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Predicate.thy Sat Jul 04 15:19:47 2009 +0200 @@ -388,10 +388,10 @@ "P \ Q = Pred (eval P \ eval Q)" definition - "\A = Pred (INFI A eval)" + [code del]: "\A = Pred (INFI A eval)" definition - "\A = Pred (SUPR A eval)" + [code del]: "\A = Pred (SUPR A eval)" instance by default (auto simp add: less_eq_pred_def less_pred_def diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Tools/Datatype/datatype_abs_proofs.ML --- a/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Sat Jul 04 15:19:47 2009 +0200 @@ -264,7 +264,7 @@ thy2 |> Sign.add_path (space_implode "_" new_type_names) |> PureThy.add_thmss [((Binding.name "recs", rec_thms), - [Nitpick_Const_Simp_Thms.add])] + [Nitpick_Const_Simps.add])] ||> Sign.parent_path ||> Theory.checkpoint |-> (fn thms => pair (reccomb_names, Library.flat thms)) @@ -337,7 +337,7 @@ (DatatypeProp.make_cases new_type_names descr sorts thy2) in thy2 - |> Context.the_theory o fold (fold Nitpick_Const_Simp_Thms.add_thm) case_thms + |> Context.the_theory o fold (fold Nitpick_Const_Simps.add_thm) case_thms o Context.Theory |> parent_path (#flat_names config) |> store_thmss "cases" new_type_names case_thms diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Tools/Function/fundef.ML --- a/src/HOL/Tools/Function/fundef.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Tools/Function/fundef.ML Sat Jul 04 15:19:47 2009 +0200 @@ -37,12 +37,12 @@ val simp_attribs = map (Attrib.internal o K) [Simplifier.simp_add, Code.add_default_eqn_attribute, - Nitpick_Const_Simp_Thms.add, - Quickcheck_RecFun_Simp_Thms.add] + Nitpick_Const_Simps.add, + Quickcheck_RecFun_Simps.add] val psimp_attribs = map (Attrib.internal o K) [Simplifier.simp_add, - Nitpick_Const_Psimp_Thms.add] + Nitpick_Const_Psimps.add] fun note_theorem ((name, atts), ths) = LocalTheory.note Thm.generatedK ((Binding.qualified_name name, atts), ths) @@ -202,7 +202,7 @@ "declaration of congruence rule for function definitions" #> setup_case_cong #> FundefRelation.setup - #> FundefCommon.TerminationSimps.setup + #> FundefCommon.Termination_Simps.setup val get_congs = FundefCtxTree.get_fundef_congs diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Tools/Function/fundef_common.ML --- a/src/HOL/Tools/Function/fundef_common.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Tools/Function/fundef_common.ML Sat Jul 04 15:19:47 2009 +0200 @@ -144,7 +144,7 @@ (* Simp rules for termination proofs *) -structure TerminationSimps = NamedThmsFun +structure Termination_Simps = Named_Thms ( val name = "termination_simp" val description = "Simplification rule for termination proofs" diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Tools/Function/lexicographic_order.ML --- a/src/HOL/Tools/Function/lexicographic_order.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Tools/Function/lexicographic_order.ML Sat Jul 04 15:19:47 2009 +0200 @@ -216,7 +216,8 @@ fun lexicographic_order_tac ctxt = TRY (FundefCommon.apply_termination_rule ctxt 1) - THEN lex_order_tac ctxt (auto_tac (local_clasimpset_of ctxt addsimps2 FundefCommon.TerminationSimps.get ctxt)) + THEN lex_order_tac ctxt + (auto_tac (local_clasimpset_of ctxt addsimps2 FundefCommon.Termination_Simps.get ctxt)) val lexicographic_order = SIMPLE_METHOD o lexicographic_order_tac diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Tools/Function/measure_functions.ML --- a/src/HOL/Tools/Function/measure_functions.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Tools/Function/measure_functions.ML Sat Jul 04 15:19:47 2009 +0200 @@ -1,7 +1,7 @@ (* Title: HOL/Tools/Function/measure_functions.ML Author: Alexander Krauss, TU Muenchen -Measure functions, generated heuristically +Measure functions, generated heuristically. *) signature MEASURE_FUNCTIONS = @@ -16,7 +16,8 @@ struct (** User-declared size functions **) -structure MeasureHeuristicRules = NamedThmsFun( +structure Measure_Heuristic_Rules = Named_Thms +( val name = "measure_function" val description = "Rules that guide the heuristic generation of measure functions" ); @@ -24,7 +25,7 @@ fun mk_is_measures t = Const (@{const_name "is_measure"}, fastype_of t --> HOLogic.boolT) $ t fun find_measures ctxt T = - DEPTH_SOLVE (resolve_tac (MeasureHeuristicRules.get ctxt) 1) + DEPTH_SOLVE (resolve_tac (Measure_Heuristic_Rules.get ctxt) 1) (HOLogic.mk_Trueprop (mk_is_measures (Var (("f",0), T --> HOLogic.natT))) |> cterm_of (ProofContext.theory_of ctxt) |> Goal.init) |> Seq.map (prop_of #> (fn _ $ (_ $ (_ $ f)) => f)) @@ -41,7 +42,7 @@ @ map (fn m => SumTree.mk_sumcase fT sT HOLogic.natT (constant_0 fT) m) (mk_funorder_funs sT) | mk_funorder_funs T = [ constant_1 T ] -fun mk_ext_base_funs ctxt (Type("+", [fT, sT])) = +fun mk_ext_base_funs ctxt (Type ("+", [fT, sT])) = map_product (SumTree.mk_sumcase fT sT HOLogic.natT) (mk_ext_base_funs ctxt fT) (mk_ext_base_funs ctxt sT) | mk_ext_base_funs ctxt T = find_measures ctxt T @@ -52,7 +53,7 @@ val get_measure_functions = mk_all_measure_funs -val setup = MeasureHeuristicRules.setup +val setup = Measure_Heuristic_Rules.setup end diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Tools/Function/scnp_reconstruct.ML --- a/src/HOL/Tools/Function/scnp_reconstruct.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Tools/Function/scnp_reconstruct.ML Sat Jul 04 15:19:47 2009 +0200 @@ -405,7 +405,7 @@ fun decomp_scnp orders ctxt = let - val extra_simps = FundefCommon.TerminationSimps.get ctxt + val extra_simps = FundefCommon.Termination_Simps.get ctxt val autom_tac = auto_tac (local_clasimpset_of ctxt addsimps2 extra_simps) in SIMPLE_METHOD diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Tools/Function/size.ML --- a/src/HOL/Tools/Function/size.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Tools/Function/size.ML Sat Jul 04 15:19:47 2009 +0200 @@ -209,7 +209,7 @@ val ([size_thms], thy'') = PureThy.add_thmss [((Binding.name "size", size_eqns), - [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add, + [Simplifier.simp_add, Nitpick_Const_Simps.add, Thm.declaration_attribute (fn thm => Context.mapping (Code.add_default_eqn thm) I)])] thy' diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Tools/arith_data.ML --- a/src/HOL/Tools/arith_data.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Tools/arith_data.ML Sat Jul 04 15:19:47 2009 +0200 @@ -28,7 +28,8 @@ (* slots for plugging in arithmetic facts and tactics *) -structure Arith_Facts = NamedThmsFun( +structure Arith_Facts = Named_Thms +( val name = "arith" val description = "arith facts - only ground formulas" ); diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Tools/inductive.ML --- a/src/HOL/Tools/inductive.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Tools/inductive.ML Sat Jul 04 15:19:47 2009 +0200 @@ -694,7 +694,7 @@ LocalTheory.notes kind (map rec_qualified intr_bindings ~~ intr_atts ~~ map (fn th => [([th], [Attrib.internal (K (ContextRules.intro_query NONE)), - Attrib.internal (K Nitpick_Ind_Intro_Thms.add)])]) intrs) |>> + Attrib.internal (K Nitpick_Ind_Intros.add)])]) intrs) |>> map (hd o snd); val (((_, elims'), (_, [induct'])), ctxt2) = ctxt1 |> diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Tools/old_primrec.ML --- a/src/HOL/Tools/old_primrec.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Tools/old_primrec.ML Sat Jul 04 15:19:47 2009 +0200 @@ -283,8 +283,8 @@ val simps'' = maps snd simps'; in thy'' - |> note (("simps", [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add, - Code.add_default_eqn_attribute]), simps'') + |> note (("simps", + [Simplifier.simp_add, Nitpick_Const_Simps.add, Code.add_default_eqn_attribute]), simps'') |> snd |> note (("induct", []), [prepare_induct (#2 (hd dts)) rec_eqns]) |> snd diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Tools/primrec.ML --- a/src/HOL/Tools/primrec.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Tools/primrec.ML Sat Jul 04 15:19:47 2009 +0200 @@ -272,7 +272,7 @@ (Binding.qualify false prefix b, Code.add_default_eqn_attrib :: attrs)) spec; fun simp_attr_binding prefix = (Binding.qualify true prefix (Binding.name "simps"), map (Attrib.internal o K) - [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add, Quickcheck_RecFun_Simp_Thms.add]); + [Simplifier.simp_add, Nitpick_Const_Simps.add, Quickcheck_RecFun_Simps.add]); in lthy |> set_group ? LocalTheory.set_group (serial_string ()) diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Tools/quickcheck_generators.ML --- a/src/HOL/Tools/quickcheck_generators.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Tools/quickcheck_generators.ML Sat Jul 04 15:19:47 2009 +0200 @@ -67,10 +67,11 @@ fun random_fun T1 T2 eq term_of random random_split seed = let - val (seed', seed'') = random_split seed; - val state = ref (seed', [], Const (@{const_name undefined}, T1 --> T2)); val fun_upd = Const (@{const_name fun_upd}, (T1 --> T2) --> T1 --> T2 --> T1 --> T2); + val (seed', seed'') = random_split seed; + + val state = ref (seed', [], fn () => Const (@{const_name undefined}, T1 --> T2)); fun random_fun' x = let val (seed, fun_map, f_t) = ! state; @@ -80,11 +81,11 @@ val t1 = term_of x; val ((y, t2), seed') = random seed; val fun_map' = (x, y) :: fun_map; - val f_t' = fun_upd $ f_t $ t1 $ t2 (); + val f_t' = fn () => fun_upd $ f_t () $ t1 $ t2 (); val _ = state := (seed', fun_map', f_t'); in y end end; - fun term_fun' () = #3 (! state); + fun term_fun' () = #3 (! state) (); in ((random_fun', term_fun'), seed'') end; @@ -244,7 +245,7 @@ |-> (fn proto_simps => prove_simps proto_simps) |-> (fn simps => LocalTheory.note Thm.generatedK ((b, Code.add_default_eqn_attrib :: map (Attrib.internal o K) - [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add, Quickcheck_RecFun_Simp_Thms.add]), + [Simplifier.simp_add, Nitpick_Const_Simps.add, Quickcheck_RecFun_Simps.add]), simps)) |> snd end diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Tools/recdef.ML --- a/src/HOL/Tools/recdef.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Tools/recdef.ML Sat Jul 04 15:19:47 2009 +0200 @@ -207,8 +207,8 @@ tfl_fn not_permissive thy cs (ss delcongs [imp_cong]) congs wfs name R eqs; val rules = (map o map) fst (partition_eq (eq_snd (op = : int * int -> bool)) rules_idx); - val simp_att = if null tcs then [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add, - Code.add_default_eqn_attribute, Quickcheck_RecFun_Simp_Thms.add] else []; + val simp_att = if null tcs then [Simplifier.simp_add, Nitpick_Const_Simps.add, + Code.add_default_eqn_attribute, Quickcheck_RecFun_Simps.add] else []; val ((simps' :: rules', [induct']), thy) = thy diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Tools/record.ML --- a/src/HOL/Tools/record.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Tools/record.ML Sat Jul 04 15:19:47 2009 +0200 @@ -2192,7 +2192,7 @@ thms_thy |> (snd oo PureThy.add_thmss) [((Binding.name "simps", sel_upd_simps), - [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add]), + [Simplifier.simp_add, Nitpick_Const_Simps.add]), ((Binding.name "iffs", iffs), [iff_add])] |> put_record name (make_record_info args parent fields extension induct_scheme') |> put_sel_upd (names @ [full_moreN]) sel_upd_simps diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Tools/res_atp.ML --- a/src/HOL/Tools/res_atp.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Tools/res_atp.ML Sat Jul 04 15:19:47 2009 +0200 @@ -219,8 +219,7 @@ handle ConstFree => false in case tm of Const ("Trueprop",_) $ (Const("op =",_) $ lhs $ rhs) => - defs lhs rhs andalso - (Output.debug (fn () => "Definition found: " ^ name ^ "_" ^ Int.toString n); true) + defs lhs rhs | _ => false end; @@ -276,8 +275,7 @@ fun relevance_filter max_new theory_const thy axioms goals = if run_relevance_filter andalso pass_mark >= 0.1 then - let val _ = Output.debug (fn () => "Start of relevance filtering"); - val const_tab = List.foldl (count_axiom_consts theory_const thy) Symtab.empty axioms + let val const_tab = List.foldl (count_axiom_consts theory_const thy) Symtab.empty axioms val goal_const_tab = get_goal_consts_typs thy goals val _ = Output.debug (fn () => ("Initial constants: " ^ space_implode ", " (Symtab.keys goal_const_tab))); @@ -406,8 +404,6 @@ fun check_named ("",th) = (warning ("No name for theorem " ^ Display.string_of_thm th); false) | check_named (_,th) = true; -fun display_thm (name,th) = Output.debug (fn () => name ^ ": " ^ Display.string_of_thm th); - (* get lemmas from claset, simpset, atpset and extra supplied rules *) fun get_clasimp_atp_lemmas ctxt = let val included_thms = @@ -419,7 +415,6 @@ let val atpset_thms = if include_atpset then ResAxioms.atpset_rules_of ctxt else [] - val _ = (Output.debug (fn () => "ATP theorems: "); app display_thm atpset_thms) in atpset_thms end in filter check_named included_thms diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Tools/res_clause.ML --- a/src/HOL/Tools/res_clause.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Tools/res_clause.ML Sat Jul 04 15:19:47 2009 +0200 @@ -381,8 +381,6 @@ | iter_type_class_pairs thy tycons classes = let val cpairs = type_class_pairs thy tycons classes val newclasses = union_all (union_all (union_all (map (map #2 o #2) cpairs))) \\ classes \\ HOLogic.typeS - val _ = if null newclasses then () - else Output.debug (fn _ => "New classes: " ^ space_implode ", " newclasses) val (classes', cpairs') = iter_type_class_pairs thy tycons newclasses in (classes' union classes, cpairs' union cpairs) end; diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/Tools/res_hol_clause.ML --- a/src/HOL/Tools/res_hol_clause.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/Tools/res_hol_clause.ML Sat Jul 04 15:19:47 2009 +0200 @@ -419,13 +419,13 @@ val ct = List.foldl (count_user_clause user_lemmas) ct0 axclauses fun needed c = valOf (Symtab.lookup ct c) > 0 val IK = if needed "c_COMBI" orelse needed "c_COMBK" - then (Output.debug (fn () => "Include combinator I K"); cnf_helper_thms thy [comb_I,comb_K]) + then cnf_helper_thms thy [comb_I,comb_K] else [] val BC = if needed "c_COMBB" orelse needed "c_COMBC" - then (Output.debug (fn () => "Include combinator B C"); cnf_helper_thms thy [comb_B,comb_C]) + then cnf_helper_thms thy [comb_B,comb_C] else [] val S = if needed "c_COMBS" - then (Output.debug (fn () => "Include combinator S"); cnf_helper_thms thy [comb_S]) + then cnf_helper_thms thy [comb_S] else [] val other = cnf_helper_thms thy [fequal_imp_equal,equal_imp_fequal] in diff -r 36156921f5e5 -r f6d9798b13f1 src/HOL/ex/Numeral.thy --- a/src/HOL/ex/Numeral.thy Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOL/ex/Numeral.thy Sat Jul 04 15:19:47 2009 +0200 @@ -93,11 +93,14 @@ subsection {* Numeral operations *} ML {* -structure DigSimps = - NamedThmsFun(val name = "numeral"; val description = "Simplification rules for numerals") +structure Dig_Simps = Named_Thms +( + val name = "numeral" + val description = "Simplification rules for numerals" +) *} -setup DigSimps.setup +setup Dig_Simps.setup instantiation num :: "{plus,times,ord}" begin diff -r 36156921f5e5 -r f6d9798b13f1 src/HOLCF/Cont.thy --- a/src/HOLCF/Cont.thy Sat Jul 04 15:19:29 2009 +0200 +++ b/src/HOLCF/Cont.thy Sat Jul 04 15:19:47 2009 +0200 @@ -140,8 +140,11 @@ subsection {* Continuity simproc *} ML {* -structure Cont2ContData = NamedThmsFun - ( val name = "cont2cont" val description = "continuity intro rule" ) +structure Cont2ContData = Named_Thms +( + val name = "cont2cont" + val description = "continuity intro rule" +) *} setup Cont2ContData.setup diff -r 36156921f5e5 -r f6d9798b13f1 src/Pure/General/symbol.scala --- a/src/Pure/General/symbol.scala Sat Jul 04 15:19:29 2009 +0200 +++ b/src/Pure/General/symbol.scala Sat Jul 04 15:19:47 2009 +0200 @@ -7,7 +7,7 @@ package isabelle import scala.io.Source -import scala.collection.jcl +import scala.collection.{jcl, mutable} import scala.util.matching.Regex @@ -34,14 +34,59 @@ def is_open(s: String): Boolean = { val len = s.length - len == 1 && Character.isHighSurrogate(s(0)) || + len == 1 && Character.isLowSurrogate(s(0)) || s == "\\" || s == "\\<" || len > 2 && s(len - 1) != '>' } - /** Recoding **/ + /** Decoding offsets **/ + + class Index(text: String) + { + case class Entry(chr: Int, sym: Int) + val index: Array[Entry] = + { + val length = text.length + val matcher = regex.pattern.matcher(text) + val buf = new mutable.ArrayBuffer[Entry] + var chr = 0 + var sym = 0 + while (chr < length) { + val c = text(chr) + val len = + if (c == '\\' || Character.isHighSurrogate(c)) { + matcher.region(chr, length).lookingAt + matcher.group.length + } else 1 + chr += len + sym += 1 + if (len > 1) buf += Entry(chr, sym) + } + buf.toArray + } + def decode(sym: Int): Int = + { + val end = index.length + def bisect(a: Int, b: Int): Int = + { + if (a < b) { + val c = (a + b) / 2 + if (sym < index(c).sym) bisect(a, c) + else if (c + 1 == end || sym < index(c + 1).sym) c + else bisect(c + 1, b) + } + else -1 + } + val i = bisect(0, end) + if (i < 0) sym + else index(i).chr + sym - index(i).sym + } + } + + + /** Recoding text **/ private class Recoder(list: List[(String, String)]) { @@ -71,8 +116,7 @@ while (i < len) { val c = text(i) if (min <= c && c <= max) { - matcher.region(i, len) - matcher.lookingAt + matcher.region(i, len).lookingAt val x = matcher.group result.append(table.get(x) getOrElse x) i = matcher.end diff -r 36156921f5e5 -r f6d9798b13f1 src/Pure/Isar/class.ML --- a/src/Pure/Isar/class.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/Pure/Isar/class.ML Sat Jul 04 15:19:47 2009 +0200 @@ -90,7 +90,7 @@ val sup_of_classes = map (snd o rules thy) sups; val loc_axiom_intros = map Drule.standard' (Locale.axioms_of thy class); val axclass_intro = #intro (AxClass.get_info thy class); - val base_sort_trivs = Drule.sort_triv thy (aT, base_sort); + val base_sort_trivs = Thm.sort_triv thy (aT, base_sort); val tac = REPEAT (SOMEGOAL (Tactic.match_tac (axclass_intro :: sup_of_classes @ loc_axiom_intros @ base_sort_trivs) diff -r 36156921f5e5 -r f6d9798b13f1 src/Pure/Proof/proof_syntax.ML --- a/src/Pure/Proof/proof_syntax.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/Pure/Proof/proof_syntax.ML Sat Jul 04 15:19:47 2009 +0200 @@ -54,6 +54,7 @@ (Binding.name "AbsP", [propT, proofT --> proofT] ---> proofT, NoSyn), (Binding.name "Hyp", propT --> proofT, NoSyn), (Binding.name "Oracle", propT --> proofT, NoSyn), + (Binding.name "Inclass", (Term.a_itselfT --> propT) --> proofT, NoSyn), (Binding.name "MinProof", proofT, Delimfix "?")] |> Sign.add_nonterminals [Binding.name "param", Binding.name "params"] |> Sign.add_syntax_i @@ -112,6 +113,12 @@ | NONE => error ("Unknown theorem " ^ quote name)) end | _ => error ("Illegal proof constant name: " ^ quote s)) + | prf_of Ts (Const ("Inclass", _) $ Const (c_class, _)) = + (case try Logic.class_of_const c_class of + SOME c => + change_type (if ty then SOME Ts else NONE) + (Inclass (TVar ((Name.aT, 0), []), c)) + | NONE => error ("Bad class constant: " ^ quote c_class)) | prf_of Ts (Const ("Hyp", _) $ prop) = Hyp prop | prf_of Ts (v as Var ((_, Type ("proof", _)))) = Hyp v | prf_of [] (Const ("Abst", _) $ Abs (s, T, prf)) = @@ -141,6 +148,7 @@ val AppPt = Const ("AppP", [proofT, proofT] ---> proofT); val Hypt = Const ("Hyp", propT --> proofT); val Oraclet = Const ("Oracle", propT --> proofT); +val Inclasst = Const ("Inclass", (Term.itselfT dummyT --> propT) --> proofT); val MinProoft = Const ("MinProof", proofT); val mk_tyapp = fold (fn T => fn prf => Const ("Appt", @@ -153,6 +161,8 @@ | term_of _ (PAxm (name, _, NONE)) = Const (Long_Name.append "axm" name, proofT) | term_of _ (PAxm (name, _, SOME Ts)) = mk_tyapp Ts (Const (Long_Name.append "axm" name, proofT)) + | term_of _ (Inclass (T, c)) = + mk_tyapp [T] (Inclasst $ Const (Logic.const_of_class c, Term.itselfT dummyT --> propT)) | term_of _ (PBound i) = Bound i | term_of Ts (Abst (s, opT, prf)) = let val T = the_default dummyT opT diff -r 36156921f5e5 -r f6d9798b13f1 src/Pure/Proof/reconstruct.ML --- a/src/Pure/Proof/reconstruct.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/Pure/Proof/reconstruct.ML Sat Jul 04 15:19:47 2009 +0200 @@ -140,7 +140,8 @@ let val tvars = OldTerm.term_tvars prop; val tfrees = OldTerm.term_tfrees prop; - val (env', Ts) = (case opTs of + val (env', Ts) = + (case opTs of NONE => Library.foldl_map mk_tvar (env, map snd tvars @ map snd tfrees) | SOME Ts => (env, Ts)); val prop' = subst_atomic_types (map TVar tvars @ map TFree tfrees ~~ Ts) @@ -222,6 +223,8 @@ mk_cnstrts_atom env vTs prop opTs prf | mk_cnstrts env _ _ vTs (prf as PAxm (_, prop, opTs)) = mk_cnstrts_atom env vTs prop opTs prf + | mk_cnstrts env _ _ vTs (prf as Inclass (T, c)) = + mk_cnstrts_atom env vTs (Logic.mk_inclass (T, c)) NONE prf | mk_cnstrts env _ _ vTs (prf as Oracle (_, prop, opTs)) = mk_cnstrts_atom env vTs prop opTs prf | mk_cnstrts env _ _ vTs (Hyp t) = (t, Hyp t, [], env, vTs) @@ -318,6 +321,7 @@ | prop_of0 Hs (Hyp t) = t | prop_of0 Hs (PThm (_, ((_, prop, SOME Ts), _))) = prop_of_atom prop Ts | prop_of0 Hs (PAxm (_, prop, SOME Ts)) = prop_of_atom prop Ts + | prop_of0 Hs (Inclass (T, c)) = Logic.mk_inclass (T, c) | prop_of0 Hs (Oracle (_, prop, SOME Ts)) = prop_of_atom prop Ts | prop_of0 _ _ = error "prop_of: partial proof object"; diff -r 36156921f5e5 -r f6d9798b13f1 src/Pure/System/gui_setup.scala --- a/src/Pure/System/gui_setup.scala Sat Jul 04 15:19:29 2009 +0200 +++ b/src/Pure/System/gui_setup.scala Sat Jul 04 15:19:47 2009 +0200 @@ -43,9 +43,6 @@ } // values - if (Platform.is_windows) { - text.append("Cygwin root: " + Cygwin.config()._1 + "\n") - } Platform.defaults match { case None => case Some((name, None)) => text.append("Platform: " + name + "\n") @@ -53,7 +50,15 @@ text.append("Main platform: " + name1 + "\n") text.append("Alternative platform: " + name2 + "\n") } - text.append("Isabelle home: " + java.lang.System.getProperty("isabelle.home")) + if (Platform.is_windows) { + text.append("Cygwin root: " + Cygwin.config()._1 + "\n") + } + try { + val isabelle_system = new Isabelle_System + text.append("Isabelle home: " + isabelle_system.getenv("ISABELLE_HOME") + "\n") + } catch { + case e: RuntimeException => text.append(e.getMessage + "\n") + } // reactions listenTo(ok) diff -r 36156921f5e5 -r f6d9798b13f1 src/Pure/System/isabelle_system.scala --- a/src/Pure/System/isabelle_system.scala Sat Jul 04 15:19:29 2009 +0200 +++ b/src/Pure/System/isabelle_system.scala Sat Jul 04 15:19:47 2009 +0200 @@ -80,18 +80,19 @@ val env0 = Map(java.lang.System.getenv.toList: _*) - val isabelle = - env0.get("ISABELLE_TOOL") match { + val isabelle_home = + env0.get("ISABELLE_HOME") match { case None | Some("") => - val isabelle = java.lang.System.getProperty("isabelle.tool") - if (isabelle == null || isabelle == "") "isabelle" - else isabelle - case Some(isabelle) => isabelle + val path = java.lang.System.getProperty("isabelle.home") + if (path == null || path == "") error("Unknown Isabelle home directory") + else path + case Some(path) => path } val dump = File.createTempFile("isabelle", null) try { - val cmdline = shell_prefix ::: List(isabelle, "getenv", "-d", dump.toString) + val cmdline = shell_prefix ::: + List(isabelle_home + "/bin/isabelle", "getenv", "-d", dump.toString) val proc = Isabelle_System.raw_execute(env0, true, cmdline: _*) val (output, rc) = Isabelle_System.process_output(proc) if (rc != 0) error(output) diff -r 36156921f5e5 -r f6d9798b13f1 src/Pure/Tools/named_thms.ML --- a/src/Pure/Tools/named_thms.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/Pure/Tools/named_thms.ML Sat Jul 04 15:19:47 2009 +0200 @@ -14,7 +14,7 @@ val setup: theory -> theory end; -functor NamedThmsFun(val name: string val description: string): NAMED_THMS = +functor Named_Thms(val name: string val description: string): NAMED_THMS = struct structure Data = GenericDataFun diff -r 36156921f5e5 -r f6d9798b13f1 src/Pure/axclass.ML --- a/src/Pure/axclass.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/Pure/axclass.ML Sat Jul 04 15:19:47 2009 +0200 @@ -580,7 +580,7 @@ | T as TVar (_, S) => insert (eq_fst op =) (T, S) | _ => I) typ []; val hyps = vars - |> map (fn (T, S) => (T, Drule.sort_triv thy (T, S) ~~ S)); + |> map (fn (T, S) => (T, Thm.sort_triv thy (T, S) ~~ S)); val ths = (typ, sort) |> Sorts.of_sort_derivation (Syntax.pp_global thy) (Sign.classes_of thy) {class_relation = diff -r 36156921f5e5 -r f6d9798b13f1 src/Pure/drule.ML --- a/src/Pure/drule.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/Pure/drule.ML Sat Jul 04 15:19:47 2009 +0200 @@ -108,7 +108,6 @@ val dummy_thm: thm val sort_constraintI: thm val sort_constraint_eq: thm - val sort_triv: theory -> typ * sort -> thm list val unconstrainTs: thm -> thm val with_subgoal: int -> (thm -> thm) -> thm -> thm val comp_no_flatten: thm * int -> int -> thm -> thm @@ -209,15 +208,6 @@ (* type classes and sorts *) -fun sort_triv thy (T, S) = - let - val certT = Thm.ctyp_of thy; - val cT = certT T; - fun class_triv c = - Thm.class_triv thy c - |> Thm.instantiate ([(certT (TVar ((Name.aT, 0), [c])), cT)], []); - in map class_triv S end; - fun unconstrainTs th = fold (Thm.unconstrainT o Thm.ctyp_of (Thm.theory_of_thm th) o TVar) (Thm.fold_terms Term.add_tvars th []) th; diff -r 36156921f5e5 -r f6d9798b13f1 src/Pure/more_thm.ML --- a/src/Pure/more_thm.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/Pure/more_thm.ML Sat Jul 04 15:19:47 2009 +0200 @@ -26,6 +26,7 @@ val eq_thm_thy: thm * thm -> bool val eq_thm_prop: thm * thm -> bool val equiv_thm: thm * thm -> bool + val sort_triv: theory -> typ * sort -> thm list val check_shyps: sort list -> thm -> thm val is_dummy: thm -> bool val plain_prop_of: thm -> term @@ -168,7 +169,16 @@ Pattern.equiv (Theory.merge (pairself Thm.theory_of_thm ths)) (pairself Thm.full_prop_of ths); -(* sort hypotheses *) +(* type classes and sorts *) + +fun sort_triv thy (T, S) = + let + val certT = Thm.ctyp_of thy; + val cT = certT T; + fun class_triv c = + Thm.class_triv thy c + |> Thm.instantiate ([(certT (TVar ((Name.aT, 0), [c])), cT)], []); + in map class_triv S end; fun check_shyps sorts raw_th = let diff -r 36156921f5e5 -r f6d9798b13f1 src/Pure/proofterm.ML --- a/src/Pure/proofterm.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/Pure/proofterm.ML Sat Jul 04 15:19:47 2009 +0200 @@ -19,6 +19,7 @@ | op %% of proof * proof | Hyp of term | PAxm of string * term * typ list option + | Inclass of typ * class | Oracle of string * term * typ list option | Promise of serial * term * typ list | PThm of serial * ((string * term * typ list option) * proof_body future) @@ -140,6 +141,7 @@ | op %% of proof * proof | Hyp of term | PAxm of string * term * typ list option + | Inclass of typ * class | Oracle of string * term * typ list option | Promise of serial * term * typ list | PThm of serial * ((string * term * typ list option) * proof_body future) @@ -290,6 +292,7 @@ | mapp (prf1 %% prf2) = (mapp prf1 %% mapph prf2 handle SAME => prf1 %% mapp prf2) | mapp (PAxm (a, prop, SOME Ts)) = PAxm (a, prop, SOME (map_typs Ts)) + | mapp (Inclass (T, c)) = Inclass (apsome g (SOME T) |> the, c) | mapp (Oracle (a, prop, SOME Ts)) = Oracle (a, prop, SOME (map_typs Ts)) | mapp (Promise (i, prop, Ts)) = Promise (i, prop, map_typs Ts) | mapp (PThm (i, ((a, prop, SOME Ts), body))) = @@ -317,6 +320,7 @@ | fold_proof_terms f g (prf1 %% prf2) = fold_proof_terms f g prf1 #> fold_proof_terms f g prf2 | fold_proof_terms _ g (PAxm (_, _, SOME Ts)) = fold g Ts + | fold_proof_terms _ g (Inclass (T, _)) = g T | fold_proof_terms _ g (Oracle (_, _, SOME Ts)) = fold g Ts | fold_proof_terms _ g (Promise (_, _, Ts)) = fold g Ts | fold_proof_terms _ g (PThm (_, ((_, _, SOME Ts), _))) = fold g Ts @@ -331,6 +335,7 @@ | size_of_proof _ = 1; fun change_type opTs (PAxm (name, prop, _)) = PAxm (name, prop, opTs) + | change_type (SOME [T]) (Inclass (_, c)) = Inclass (T, c) | change_type opTs (Oracle (name, prop, _)) = Oracle (name, prop, opTs) | change_type opTs (Promise _) = error "change_type: unexpected promise" | change_type opTs (PThm (i, ((name, prop, _), body))) = PThm (i, ((name, prop, opTs), body)) @@ -468,6 +473,7 @@ | norm (prf1 %% prf2) = (norm prf1 %% normh prf2 handle SAME => prf1 %% norm prf2) | norm (PAxm (s, prop, Ts)) = PAxm (s, prop, apsome' (htypeTs norm_types_same) Ts) + | norm (Inclass (T, c)) = Inclass (htypeT norm_type_same T, c) | norm (Oracle (s, prop, Ts)) = Oracle (s, prop, apsome' (htypeTs norm_types_same) Ts) | norm (Promise (i, prop, Ts)) = Promise (i, prop, htypeTs norm_types_same Ts) | norm (PThm (i, ((s, t, Ts), body))) = @@ -713,6 +719,8 @@ handle SAME => prf1 %% lift' Us Ts prf2) | lift' _ _ (PAxm (s, prop, Ts)) = PAxm (s, prop, apsome' (same (op =) (map (Logic.incr_tvar inc))) Ts) + | lift' _ _ (Inclass (T, c)) = + Inclass (same (op =) (Logic.incr_tvar inc) T, c) | lift' _ _ (Oracle (s, prop, Ts)) = Oracle (s, prop, apsome' (same (op =) (map (Logic.incr_tvar inc))) Ts) | lift' _ _ (Promise (i, prop, Ts)) = @@ -967,8 +975,9 @@ orelse has_duplicates (op =) (Library.foldl (fn (js, SOME (Bound j)) => j :: js | (js, _) => js) ([], ts)) orelse exists #1 prfs then [i] else [], false, map (pair false) ts, prf) - | shrink' ls lev ts prfs (Hyp t) = ([], false, map (pair false) ts, Hyp t) - | shrink' ls lev ts prfs MinProof = ([], false, map (pair false) ts, MinProof) + | shrink' ls lev ts prfs (prf as Hyp _) = ([], false, map (pair false) ts, prf) + | shrink' ls lev ts prfs (prf as MinProof) = ([], false, map (pair false) ts, prf) + | shrink' ls lev ts prfs (prf as Inclass _) = ([], false, map (pair false) ts, prf) | shrink' ls lev ts prfs prf = let val prop = @@ -1060,6 +1069,9 @@ | mtch Ts i j inst (PAxm (s1, _, opTs), PAxm (s2, _, opUs)) = if s1 = s2 then optmatch matchTs inst (opTs, opUs) else raise PMatch + | mtch Ts i j inst (Inclass (T1, c1), Inclass (T2, c2)) = + if c1 = c2 then matchT inst (T1, T2) + else raise PMatch | mtch Ts i j inst (PThm (_, ((name1, prop1, opTs), _)), PThm (_, ((name2, prop2, opUs), _))) = if name1 = name2 andalso prop1 = prop2 then optmatch matchTs inst (opTs, opUs) @@ -1091,6 +1103,7 @@ NONE => prf | SOME prf' => incr_pboundvars plev tlev prf') | subst _ _ (PAxm (id, prop, Ts)) = PAxm (id, prop, Option.map (map substT) Ts) + | subst _ _ (Inclass (T, c)) = Inclass (substT T, c) | subst _ _ (Oracle (id, prop, Ts)) = Oracle (id, prop, Option.map (map substT) Ts) | subst _ _ (Promise (i, prop, Ts)) = Promise (i, prop, (map substT) Ts) | subst _ _ (PThm (i, ((id, prop, Ts), body))) = @@ -1117,6 +1130,7 @@ (_, Hyp (Var _)) => true | (Hyp (Var _), _) => true | (PAxm (a, _, _), PAxm (b, _, _)) => a = b andalso matchrands prf1 prf2 + | (Inclass (_, c), Inclass (_, d)) => c = d andalso matchrands prf1 prf2 | (PThm (_, ((a, propa, _), _)), PThm (_, ((b, propb, _), _))) => a = b andalso propa = propb andalso matchrands prf1 prf2 | (PBound i, PBound j) => i = j andalso matchrands prf1 prf2 diff -r 36156921f5e5 -r f6d9798b13f1 src/Pure/thm.ML --- a/src/Pure/thm.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/Pure/thm.ML Sat Jul 04 15:19:47 2009 +0200 @@ -487,7 +487,8 @@ val extra' = Sorts.subtract (map #2 (Sign.witness_sorts thy present extra)) extra |> Sorts.minimal_sorts (Sign.classes_of thy); - val shyps' = Sorts.union present extra'; + val shyps' = Sorts.union present extra' + |> Sorts.remove_sort []; in Thm (der, {thy_ref = Theory.check_thy thy, tags = tags, maxidx = maxidx, shyps = shyps', hyps = hyps, tpairs = tpairs, prop = prop}) @@ -1110,15 +1111,21 @@ prop = Logic.mk_implies (A, A)}); (*Axiom-scheme reflecting signature contents: "OFCLASS(?'a::c, c_class)" *) -fun class_triv thy c = +fun class_triv thy raw_c = let - val Cterm {t, maxidx, sorts, ...} = - cterm_of thy (Logic.mk_inclass (TVar ((Name.aT, 0), [c]), Sign.certify_class thy c)) - handle TERM (msg, _) => raise THM ("class_triv: " ^ msg, 0, []); - val der = deriv_rule0 (Pt.PAxm ("Pure.class_triv:" ^ c, t, SOME [])); + val c = Sign.certify_class thy raw_c; + val T = TVar ((Name.aT, 0), [c]); + val Cterm {t = prop, maxidx, sorts, ...} = cterm_of thy (Logic.mk_inclass (T, c)) + handle TERM (msg, _) => raise THM ("class_triv: " ^ msg, 0, []); in - Thm (der, {thy_ref = Theory.check_thy thy, tags = [], maxidx = maxidx, - shyps = sorts, hyps = [], tpairs = [], prop = t}) + Thm (deriv_rule0 (Pt.Inclass (T, c)), + {thy_ref = Theory.check_thy thy, + tags = [], + maxidx = maxidx, + shyps = sorts, + hyps = [], + tpairs = [], + prop = prop}) end; (*Internalize sort constraints of type variable*) diff -r 36156921f5e5 -r f6d9798b13f1 src/Tools/Code/code_haskell.ML --- a/src/Tools/Code/code_haskell.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/Tools/Code/code_haskell.ML Sat Jul 04 15:19:47 2009 +0200 @@ -147,10 +147,10 @@ val consts = map_filter (fn c => if (is_some o syntax_const) c then NONE else (SOME o Long_Name.base_name o deresolve) c) - ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []); + (fold Code_Thingol.add_constnames (t :: ts) []); val vars = init_syms |> Code_Printer.intro_vars consts - |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames) + |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_varnames) (insert (op =)) ts []); in semicolon ( diff -r 36156921f5e5 -r f6d9798b13f1 src/Tools/Code/code_ml.ML --- a/src/Tools/Code/code_ml.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/Tools/Code/code_ml.ML Sat Jul 04 15:19:47 2009 +0200 @@ -178,7 +178,7 @@ val consts = map_filter (fn c => if (is_some o syntax_const) c then NONE else (SOME o Long_Name.base_name o deresolve) c) - (Code_Thingol.fold_constnames (insert (op =)) t []); + (Code_Thingol.add_constnames t []); val vars = reserved_names |> Code_Printer.intro_vars consts; in @@ -203,10 +203,10 @@ val consts = map_filter (fn c => if (is_some o syntax_const) c then NONE else (SOME o Long_Name.base_name o deresolve) c) - ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []); + (fold Code_Thingol.add_constnames (t :: ts) []); val vars = reserved_names |> Code_Printer.intro_vars consts - |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames) + |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_varnames) (insert (op =)) ts []); in concat ( @@ -488,7 +488,7 @@ val consts = map_filter (fn c => if (is_some o syntax_const) c then NONE else (SOME o Long_Name.base_name o deresolve) c) - (Code_Thingol.fold_constnames (insert (op =)) t []); + (Code_Thingol.add_constnames t []); val vars = reserved_names |> Code_Printer.intro_vars consts; in @@ -508,10 +508,10 @@ val consts = map_filter (fn c => if (is_some o syntax_const) c then NONE else (SOME o Long_Name.base_name o deresolve) c) - ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []); + (fold Code_Thingol.add_constnames (t :: ts) []); val vars = reserved_names |> Code_Printer.intro_vars consts - |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames) + |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_varnames) (insert (op =)) ts []); in concat [ (Pretty.block o Pretty.commas) @@ -524,10 +524,10 @@ val consts = map_filter (fn c => if (is_some o syntax_const) c then NONE else (SOME o Long_Name.base_name o deresolve) c) - ((fold o Code_Thingol.fold_constnames) (insert (op =)) (t :: ts) []); + (fold Code_Thingol.add_constnames (t :: ts) []); val vars = reserved_names |> Code_Printer.intro_vars consts - |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_unbound_varnames) + |> Code_Printer.intro_vars ((fold o Code_Thingol.fold_varnames) (insert (op =)) ts []); in concat ( @@ -552,8 +552,7 @@ val consts = map_filter (fn c => if (is_some o syntax_const) c then NONE else (SOME o Long_Name.base_name o deresolve) c) - ((fold o Code_Thingol.fold_constnames) - (insert (op =)) (map (snd o fst) eqs) []); + (fold Code_Thingol.add_constnames (map (snd o fst) eqs) []); val vars = reserved_names |> Code_Printer.intro_vars consts; val dummy_parms = (map str o fish_params vars o map (fst o fst)) eqs; @@ -777,8 +776,7 @@ val (eqs', is_value) = if null (filter_out (null o snd) vs) then case eqs of [(([], t), thm)] => if (not o null o fst o Code_Thingol.unfold_fun) ty then ([(([IVar (SOME "x")], t `$ IVar (SOME "x")), thm)], false) - else (eqs, not (Code_Thingol.fold_constnames - (fn name' => fn b => b orelse name = name') t false)) + else (eqs, not (member (op =) (Code_Thingol.add_constnames t []) name)) | _ => (eqs, false) else (eqs, false) in ((name, (tysm, eqs')), is_value) end; diff -r 36156921f5e5 -r f6d9798b13f1 src/Tools/Code/code_thingol.ML --- a/src/Tools/Code/code_thingol.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/Tools/Code/code_thingol.ML Sat Jul 04 15:19:47 2009 +0200 @@ -49,9 +49,8 @@ val eta_expand: int -> const * iterm list -> iterm val contains_dictvar: iterm -> bool val locally_monomorphic: iterm -> bool - val fold_constnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a + val add_constnames: iterm -> string list -> string list val fold_varnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a - val fold_unbound_varnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a type naming val empty_naming: naming @@ -153,38 +152,30 @@ of (IConst c, ts) => SOME (c, ts) | _ => NONE; -fun fold_aiterms f (t as IConst _) = f t - | fold_aiterms f (t as IVar _) = f t - | fold_aiterms f (t1 `$ t2) = fold_aiterms f t1 #> fold_aiterms f t2 - | fold_aiterms f (t as _ `|=> t') = f t #> fold_aiterms f t' - | fold_aiterms f (ICase (_, t)) = fold_aiterms f t; - -fun fold_constnames f = - let - fun add (IConst (c, _)) = f c - | add _ = I; - in fold_aiterms add end; +fun add_constnames (IConst (c, _)) = insert (op =) c + | add_constnames (IVar _) = I + | add_constnames (t1 `$ t2) = add_constnames t1 #> add_constnames t2 + | add_constnames (_ `|=> t) = add_constnames t + | add_constnames (ICase (((t, _), ds), _)) = add_constnames t + #> fold (fn (pat, body) => add_constnames pat #> add_constnames body) ds; fun fold_varnames f = let - fun add (IVar (SOME v)) = f v - | add ((SOME v, _) `|=> _) = f v - | add _ = I; - in fold_aiterms add end; + fun fold_aux add f = + let + fun fold_term _ (IConst _) = I + | fold_term vs (IVar (SOME v)) = if member (op =) vs v then I else f v + | fold_term _ (IVar NONE) = I + | fold_term vs (t1 `$ t2) = fold_term vs t1 #> fold_term vs t2 + | fold_term vs ((SOME v, _) `|=> t) = fold_term (insert (op =) v vs) t + | fold_term vs ((NONE, _) `|=> t) = fold_term vs t + | fold_term vs (ICase (((t, _), ds), _)) = fold_term vs t #> fold (fold_case vs) ds + and fold_case vs (p, t) = fold_term (add p vs) t; + in fold_term [] end; + fun add t = fold_aux add (insert (op =)) t; + in fold_aux add f end; -fun fold_unbound_varnames f = - let - fun add _ (IConst _) = I - | add vs (IVar (SOME v)) = if not (member (op =) vs v) then f v else I - | add _ (IVar NONE) = I - | add vs (t1 `$ t2) = add vs t1 #> add vs t2 - | add vs ((SOME v, _) `|=> t) = add (insert (op =) v vs) t - | add vs ((NONE, _) `|=> t) = add vs t - | add vs (ICase (((t, _), ds), _)) = add vs t #> fold (add_case vs) ds - and add_case vs (p, t) = add (fold_varnames (insert (op =)) p vs) t; - in add [] end; - -fun exists_var t v = fold_unbound_varnames (fn w => fn b => v = w orelse b) t false; +fun exists_var t v = fold_varnames (fn w => fn b => v = w orelse b) t false; fun split_pat_abs ((NONE, ty) `|=> t) = SOME ((IVar NONE, ty), t) | split_pat_abs ((SOME v, ty) `|=> t) = SOME (case t @@ -219,12 +210,14 @@ fun contains_dictvar t = let - fun contains (DictConst (_, dss)) = (fold o fold) contains dss - | contains (DictVar _) = K true; - in - fold_aiterms - (fn IConst (_, ((_, dss), _)) => (fold o fold) contains dss | _ => I) t false - end; + fun cont_dict (DictConst (_, dss)) = (exists o exists) cont_dict dss + | cont_dict (DictVar _) = true; + fun cont_term (IConst (_, ((_, dss), _))) = (exists o exists) cont_dict dss + | cont_term (IVar _) = false + | cont_term (t1 `$ t2) = cont_term t1 orelse cont_term t2 + | cont_term (_ `|=> t) = cont_term t + | cont_term (ICase (_, t)) = cont_term t; + in cont_term t end; fun locally_monomorphic (IConst _) = false | locally_monomorphic (IVar _) = true @@ -640,20 +633,37 @@ else map2 mk_constr case_pats (nth_drop t_pos ts); fun casify naming constrs ty ts = let + val undefineds = map_filter (lookup_const naming) (Code.undefineds thy); + fun collapse_clause vs_map ts body = + let + in case body + of IConst (c, _) => if member (op =) undefineds c + then [] + else [(ts, body)] + | ICase (((IVar (SOME v), _), subclauses), _) => + if forall (fn (pat', body') => exists_var pat' v + orelse not (exists_var body' v)) subclauses + then case AList.lookup (op =) vs_map v + of SOME i => maps (fn (pat', body') => + collapse_clause (AList.delete (op =) v vs_map) + (nth_map i (K pat') ts) body') subclauses + | NONE => [(ts, body)] + else [(ts, body)] + | _ => [(ts, body)] + end; + fun mk_clause mk tys t = + let + val (vs, body) = unfold_abs_eta tys t; + val vs_map = fold_index (fn (i, (SOME v, _)) => cons (v, i) | _ => I) vs []; + val ts = map (IVar o fst) vs; + in map mk (collapse_clause vs_map ts body) end; val t = nth ts t_pos; val ts_clause = nth_drop t_pos ts; - val undefineds = map_filter (lookup_const naming) (Code.undefineds thy); - fun mk_clause ((constr as IConst (_, (_, tys)), n), t) = - let - val (vs, t') = unfold_abs_eta (curry Library.take n tys) t; - val is_undefined = case t - of IConst (c, _) => member (op =) undefineds c - | _ => false; - in if is_undefined then NONE else SOME (constr `$$ map (IVar o fst) vs, t') end; - val clauses = if null case_pats then let val ([(v, _)], t) = - unfold_abs_eta [ty] (the_single ts_clause) - in [(IVar v, t)] end - else map_filter mk_clause (constrs ~~ ts_clause); + val clauses = if null case_pats + then mk_clause (fn ([t], body) => (t, body)) [ty] (the_single ts_clause) + else maps (fn ((constr as IConst (_, (_, tys)), n), t) => + mk_clause (fn (ts, body) => (constr `$$ ts, body)) (curry Library.take n tys) t) + (constrs ~~ ts_clause); in ((t, ty), clauses) end; in translate_const thy algbr funcgr thm c_ty diff -r 36156921f5e5 -r f6d9798b13f1 src/Tools/atomize_elim.ML --- a/src/Tools/atomize_elim.ML Sat Jul 04 15:19:29 2009 +0200 +++ b/src/Tools/atomize_elim.ML Sat Jul 04 15:19:47 2009 +0200 @@ -20,7 +20,9 @@ struct (* theory data *) -structure AtomizeElimData = NamedThmsFun( + +structure AtomizeElimData = Named_Thms +( val name = "atomize_elim"; val description = "atomize_elim rewrite rule"; ); diff -r 36156921f5e5 -r f6d9798b13f1 src/ZF/UNITY/Constrains.thy --- a/src/ZF/UNITY/Constrains.thy Sat Jul 04 15:19:29 2009 +0200 +++ b/src/ZF/UNITY/Constrains.thy Sat Jul 04 15:19:47 2009 +0200 @@ -462,7 +462,11 @@ val Always_Int_rule = foldr1 (fn (th1,th2) => [th1,th2] MRS @{thm Always_Int_I}); (*To allow expansion of the program's definition when appropriate*) -structure ProgramDefs = NamedThmsFun(val name = "program" val description = "program definitions"); +structure Program_Defs = Named_Thms +( + val name = "program" + val description = "program definitions" +); (*proves "co" properties when the program is specified*) @@ -478,7 +482,7 @@ (* Three subgoals *) rewrite_goal_tac [@{thm st_set_def}] 3, REPEAT (force_tac css 2), - full_simp_tac (ss addsimps (ProgramDefs.get ctxt)) 1, + full_simp_tac (ss addsimps (Program_Defs.get ctxt)) 1, ALLGOALS (clarify_tac cs), REPEAT (FIRSTGOAL (etac disjE)), ALLGOALS (clarify_tac cs), @@ -493,7 +497,7 @@ rtac @{thm AlwaysI} i THEN force_tac (local_clasimpset_of ctxt) i THEN constrains_tac ctxt i; *} -setup ProgramDefs.setup +setup Program_Defs.setup method_setup safety = {* Scan.succeed (SIMPLE_METHOD' o constrains_tac) *} diff -r 36156921f5e5 -r f6d9798b13f1 src/ZF/UNITY/SubstAx.thy --- a/src/ZF/UNITY/SubstAx.thy Sat Jul 04 15:19:29 2009 +0200 +++ b/src/ZF/UNITY/SubstAx.thy Sat Jul 04 15:19:47 2009 +0200 @@ -359,7 +359,7 @@ REPEAT (ares_tac [@{thm LeadsTo_Basis}, @{thm leadsTo_Basis}, @{thm EnsuresI}, @{thm ensuresI}] 1), (*now there are two subgoals: co & transient*) - simp_tac (ss addsimps (ProgramDefs.get ctxt)) 2, + simp_tac (ss addsimps (Program_Defs.get ctxt)) 2, res_inst_tac ctxt [(("act", 0), sact)] @{thm transientI} 2, (*simplify the command's domain*) simp_tac (ss addsimps [@{thm domain_def}]) 3,