--- a/Admin/isatest/isatest-makedist Thu Mar 25 17:55:55 2010 +0100
+++ b/Admin/isatest/isatest-makedist Thu Mar 25 17:56:31 2010 +0100
@@ -96,7 +96,7 @@
sleep 15
$SSH macbroy22 "$MAKEALL $HOME/settings/at-poly"
sleep 15
-$SSH macbroy23 "$MAKEALL -l HOL HOL-ex $HOME/settings/at-sml-dev-e"
+$SSH macbroy23 "$MAKEALL -l HOL HOL-Library $HOME/settings/at-sml-dev-e"
sleep 15
$SSH macbroy24 "$MAKEALL $HOME/settings/at64-poly"
sleep 15
--- a/doc-src/IsarImplementation/Thy/Logic.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/doc-src/IsarImplementation/Thy/Logic.thy Thu Mar 25 17:56:31 2010 +0100
@@ -552,14 +552,13 @@
@{index_ML Thm.implies_elim: "thm -> thm -> thm"} \\
@{index_ML Thm.generalize: "string list * string list -> int -> thm -> thm"} \\
@{index_ML Thm.instantiate: "(ctyp * ctyp) list * (cterm * cterm) list -> thm -> thm"} \\
- @{index_ML Thm.axiom: "theory -> string -> thm"} \\
+ @{index_ML Thm.add_axiom: "binding * term -> theory -> thm * theory"} \\
@{index_ML Thm.add_oracle: "binding * ('a -> cterm) -> theory
-> (string * ('a -> thm)) * theory"} \\
+ @{index_ML Thm.add_def: "bool -> bool -> binding * term -> theory -> thm * theory"} \\
\end{mldecls}
\begin{mldecls}
- @{index_ML Theory.add_axioms_i: "(binding * term) list -> theory -> theory"} \\
@{index_ML Theory.add_deps: "string -> string * typ -> (string * typ) list -> theory -> theory"} \\
- @{index_ML Theory.add_defs_i: "bool -> bool -> (binding * term) list -> theory -> theory"} \\
\end{mldecls}
\begin{description}
@@ -607,26 +606,28 @@
term variables. Note that the types in @{text "\<^vec>x\<^isub>\<tau>"}
refer to the instantiated versions.
- \item @{ML Thm.axiom}~@{text "thy name"} retrieves a named
- axiom, cf.\ @{text "axiom"} in \figref{fig:prim-rules}.
+ \item @{ML Thm.add_axiom}~@{text "(name, A) thy"} declares an
+ arbitrary proposition as axiom, and retrieves it as a theorem from
+ the resulting theory, cf.\ @{text "axiom"} in
+ \figref{fig:prim-rules}. Note that the low-level representation in
+ the axiom table may differ slightly from the returned theorem.
\item @{ML Thm.add_oracle}~@{text "(binding, oracle)"} produces a named
oracle rule, essentially generating arbitrary axioms on the fly,
cf.\ @{text "axiom"} in \figref{fig:prim-rules}.
- \item @{ML Theory.add_axioms_i}~@{text "[(name, A), \<dots>]"} declares
- arbitrary propositions as axioms.
+ \item @{ML Thm.add_def}~@{text "unchecked overloaded (name, c
+ \<^vec>x \<equiv> t)"} states a definitional axiom for an existing constant
+ @{text "c"}. Dependencies are recorded via @{ML Theory.add_deps},
+ unless the @{text "unchecked"} option is set. Note that the
+ low-level representation in the axiom table may differ slightly from
+ the returned theorem.
\item @{ML Theory.add_deps}~@{text "name c\<^isub>\<tau>
\<^vec>d\<^isub>\<sigma>"} declares dependencies of a named specification
for constant @{text "c\<^isub>\<tau>"}, relative to existing
specifications for constants @{text "\<^vec>d\<^isub>\<sigma>"}.
- \item @{ML Theory.add_defs_i}~@{text "unchecked overloaded [(name, c
- \<^vec>x \<equiv> t), \<dots>]"} states a definitional axiom for an existing
- constant @{text "c"}. Dependencies are recorded (cf.\ @{ML
- Theory.add_deps}), unless the @{text "unchecked"} option is set.
-
\end{description}
*}
--- a/doc-src/IsarImplementation/Thy/document/Logic.tex Thu Mar 25 17:55:55 2010 +0100
+++ b/doc-src/IsarImplementation/Thy/document/Logic.tex Thu Mar 25 17:56:31 2010 +0100
@@ -560,14 +560,13 @@
\indexdef{}{ML}{Thm.implies\_elim}\verb|Thm.implies_elim: thm -> thm -> thm| \\
\indexdef{}{ML}{Thm.generalize}\verb|Thm.generalize: string list * string list -> int -> thm -> thm| \\
\indexdef{}{ML}{Thm.instantiate}\verb|Thm.instantiate: (ctyp * ctyp) list * (cterm * cterm) list -> thm -> thm| \\
- \indexdef{}{ML}{Thm.axiom}\verb|Thm.axiom: theory -> string -> thm| \\
+ \indexdef{}{ML}{Thm.add\_axiom}\verb|Thm.add_axiom: binding * term -> theory -> thm * theory| \\
\indexdef{}{ML}{Thm.add\_oracle}\verb|Thm.add_oracle: binding * ('a -> cterm) -> theory|\isasep\isanewline%
\verb| -> (string * ('a -> thm)) * theory| \\
+ \indexdef{}{ML}{Thm.add\_def}\verb|Thm.add_def: bool -> bool -> binding * term -> theory -> thm * theory| \\
\end{mldecls}
\begin{mldecls}
- \indexdef{}{ML}{Theory.add\_axioms\_i}\verb|Theory.add_axioms_i: (binding * term) list -> theory -> theory| \\
\indexdef{}{ML}{Theory.add\_deps}\verb|Theory.add_deps: string -> string * typ -> (string * typ) list -> theory -> theory| \\
- \indexdef{}{ML}{Theory.add\_defs\_i}\verb|Theory.add_defs_i: bool -> bool -> (binding * term) list -> theory -> theory| \\
\end{mldecls}
\begin{description}
@@ -612,23 +611,26 @@
term variables. Note that the types in \isa{\isactrlvec x\isactrlisub {\isasymtau}}
refer to the instantiated versions.
- \item \verb|Thm.axiom|~\isa{thy\ name} retrieves a named
- axiom, cf.\ \isa{axiom} in \figref{fig:prim-rules}.
+ \item \verb|Thm.add_axiom|~\isa{{\isacharparenleft}name{\isacharcomma}\ A{\isacharparenright}\ thy} declares an
+ arbitrary proposition as axiom, and retrieves it as a theorem from
+ the resulting theory, cf.\ \isa{axiom} in
+ \figref{fig:prim-rules}. Note that the low-level representation in
+ the axiom table may differ slightly from the returned theorem.
\item \verb|Thm.add_oracle|~\isa{{\isacharparenleft}binding{\isacharcomma}\ oracle{\isacharparenright}} produces a named
oracle rule, essentially generating arbitrary axioms on the fly,
cf.\ \isa{axiom} in \figref{fig:prim-rules}.
- \item \verb|Theory.add_axioms_i|~\isa{{\isacharbrackleft}{\isacharparenleft}name{\isacharcomma}\ A{\isacharparenright}{\isacharcomma}\ {\isasymdots}{\isacharbrackright}} declares
- arbitrary propositions as axioms.
+ \item \verb|Thm.add_def|~\isa{unchecked\ overloaded\ {\isacharparenleft}name{\isacharcomma}\ c\ \isactrlvec x\ {\isasymequiv}\ t{\isacharparenright}} states a definitional axiom for an existing constant
+ \isa{c}. Dependencies are recorded via \verb|Theory.add_deps|,
+ unless the \isa{unchecked} option is set. Note that the
+ low-level representation in the axiom table may differ slightly from
+ the returned theorem.
\item \verb|Theory.add_deps|~\isa{name\ c\isactrlisub {\isasymtau}\ \isactrlvec d\isactrlisub {\isasymsigma}} declares dependencies of a named specification
for constant \isa{c\isactrlisub {\isasymtau}}, relative to existing
specifications for constants \isa{\isactrlvec d\isactrlisub {\isasymsigma}}.
- \item \verb|Theory.add_defs_i|~\isa{unchecked\ overloaded\ {\isacharbrackleft}{\isacharparenleft}name{\isacharcomma}\ c\ \isactrlvec x\ {\isasymequiv}\ t{\isacharparenright}{\isacharcomma}\ {\isasymdots}{\isacharbrackright}} states a definitional axiom for an existing
- constant \isa{c}. Dependencies are recorded (cf.\ \verb|Theory.add_deps|), unless the \isa{unchecked} option is set.
-
\end{description}%
\end{isamarkuptext}%
\isamarkuptrue%
--- a/etc/components Thu Mar 25 17:55:55 2010 +0100
+++ b/etc/components Thu Mar 25 17:56:31 2010 +0100
@@ -13,7 +13,6 @@
#misc components
src/Tools/Code
src/Tools/WWW_Find
-src/Tools/Cache_IO
src/HOL/Tools/ATP_Manager
src/HOL/Mirabelle
src/HOL/Library/Sum_Of_Squares
--- a/src/HOL/Big_Operators.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOL/Big_Operators.thy Thu Mar 25 17:56:31 2010 +0100
@@ -804,7 +804,7 @@
definition (in comm_monoid_mult) setprod :: "('b \<Rightarrow> 'a) => 'b set => 'a" where
"setprod f A = (if finite A then fold_image (op *) f 1 A else 1)"
-sublocale comm_monoid_add < setprod!: comm_monoid_big "op *" 1 setprod proof
+sublocale comm_monoid_mult < setprod!: comm_monoid_big "op *" 1 setprod proof
qed (fact setprod_def)
abbreviation
--- a/src/HOL/Boogie/Examples/Boogie_Dijkstra.certs Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOL/Boogie/Examples/Boogie_Dijkstra.certs Thu Mar 25 17:56:31 2010 +0100
@@ -1,4 +1,4 @@
-JinTdmjIiorL0/vvOyf3+w 6542 0
+6164a2366a9e57e212b8ac4aa01e3c8bcc1ea8e0 6542 0
#2 := false
decl up_6 :: (-> T4 T2 bool)
decl ?x47!7 :: (-> T2 T2)
--- a/src/HOL/Boogie/Examples/Boogie_Max.certs Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOL/Boogie/Examples/Boogie_Max.certs Thu Mar 25 17:56:31 2010 +0100
@@ -1,4 +1,4 @@
-iks4GfP7O/NgNFyGZ4ynjQ 2224 0
+cdfef9f27f2a4ba9648f86890c8563d0a1cfe888 2224 0
#2 := false
#4 := 0::int
decl uf_3 :: (-> int int)
--- a/src/HOL/Boogie/Examples/VCC_Max.certs Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOL/Boogie/Examples/VCC_Max.certs Thu Mar 25 17:56:31 2010 +0100
@@ -1,4 +1,4 @@
-6Q8QWdFv463DpfVfkr0XnA 7790 0
+b95bf7adc1e2b959cf13db317b64554768249b2e 7790 0
#2 := false
decl uf_110 :: (-> T4 T5 int)
decl uf_66 :: (-> T5 int T3 T5)
--- a/src/HOL/Boogie/Tools/boogie_loader.ML Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOL/Boogie/Tools/boogie_loader.ML Thu Mar 25 17:56:31 2010 +0100
@@ -240,13 +240,10 @@
fun add_axioms verbose axs thy =
let
val (used, new) = mark_axioms thy (name_axioms axs)
- fun add (n, t) =
- Specification.axiomatization [] [((Binding.name n, []), [t])] #>>
- hd o hd o snd
in
thy
- |> fold_map add new
- |-> Context.theory_map o fold Boogie_Axioms.add_thm
+ |> fold_map (fn (n, t) => Specification.axiom ((Binding.name n, []), t)) new
+ |-> Context.theory_map o fold (Boogie_Axioms.add_thm o Drule.export_without_context)
|> log verbose "The following axioms were added:" (map fst new)
|> (fn thy' => log verbose "The following axioms already existed:"
(map (Display.string_of_thm_global thy') used) thy')
--- a/src/HOL/IsaMakefile Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOL/IsaMakefile Thu Mar 25 17:56:31 2010 +0100
@@ -10,13 +10,14 @@
images: \
HOL \
HOL-Algebra \
+ HOL-Base \
HOL-Boogie \
- HOL-Base \
HOL-Main \
HOL-Multivariate_Analysis \
HOL-NSA \
HOL-Nominal \
HOL-Plain \
+ HOL-Probability \
HOL-Proofs \
HOL-SMT \
HOL-Word \
@@ -54,7 +55,7 @@
HOL-Number_Theory \
HOL-Old_Number_Theory \
HOL-Quotient_Examples \
- HOL-Probability \
+ HOL-Predicate_Compile_Examples \
HOL-Prolog \
HOL-Proofs-Extraction \
HOL-Proofs-Lambda \
@@ -419,7 +420,8 @@
Library/Nat_Bijection.thy $(SRC)/Tools/float.ML \
$(SRC)/HOL/Tools/float_arith.ML Library/positivstellensatz.ML \
Library/reify_data.ML Library/reflection.ML Library/LaTeXsugar.thy \
- Library/OptionalSugar.thy Library/SML_Quickcheck.thy
+ Library/OptionalSugar.thy \
+ Library/Predicate_Compile_Quickcheck.thy Library/SML_Quickcheck.thy
@cd Library; $(ISABELLE_TOOL) usedir $(OUT)/HOL Library
@@ -949,7 +951,6 @@
$(LOG)/HOL-ex.gz: $(OUT)/HOL Decision_Procs/Commutative_Ring.thy \
Number_Theory/Primes.thy \
- Tools/Predicate_Compile/predicate_compile_core.ML \
ex/Abstract_NAT.thy ex/Antiquote.thy ex/Arith_Examples.thy \
ex/Arithmetic_Series_Complex.thy ex/BT.thy ex/BinEx.thy \
ex/Binary.thy ex/CTL.thy ex/Chinese.thy ex/Classical.thy \
@@ -965,7 +966,6 @@
ex/Lagrange.thy ex/LocaleTest2.thy ex/MT.thy \
ex/MergeSort.thy ex/Meson_Test.thy ex/MonoidGroup.thy \
ex/Multiquote.thy ex/NatSum.thy ex/Numeral.thy ex/PER.thy \
- ex/Predicate_Compile_ex.thy ex/Predicate_Compile_Quickcheck.thy \
ex/PresburgerEx.thy ex/Primrec.thy ex/Quickcheck_Examples.thy \
ex/ROOT.ML ex/Recdefs.thy ex/Records.thy ex/ReflectionEx.thy \
ex/Refute_Examples.thy ex/SAT_Examples.thy ex/SVC_Oracle.thy \
@@ -1079,25 +1079,24 @@
Multivariate_Analysis/Topology_Euclidean_Space.thy \
Multivariate_Analysis/Convex_Euclidean_Space.thy \
Multivariate_Analysis/Brouwer_Fixpoint.thy \
- Multivariate_Analysis/Derivative.thy
+ Multivariate_Analysis/Derivative.thy \
+ Multivariate_Analysis/Integration.thy \
+ Multivariate_Analysis/Integration.cert \
+ Multivariate_Analysis/Real_Integration.thy
@cd Multivariate_Analysis; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL-SMT HOL-Multivariate_Analysis
## HOL-Probability
-HOL-Probability: HOL $(LOG)/HOL-Probability.gz
+HOL-Probability: HOL $(OUT)/HOL-Probability
-$(LOG)/HOL-Probability.gz: $(OUT)/HOL Probability/ROOT.ML \
- Probability/Probability.thy \
- Probability/Sigma_Algebra.thy \
- Probability/SeriesPlus.thy \
- Probability/Caratheodory.thy \
- Probability/Borel.thy \
- Probability/Measure.thy \
- Probability/Lebesgue.thy \
- Probability/Product_Measure.thy \
+$(OUT)/HOL-Probability: $(OUT)/HOL Probability/ROOT.ML \
+ Probability/Probability.thy Probability/Sigma_Algebra.thy \
+ Probability/SeriesPlus.thy Probability/Caratheodory.thy \
+ Probability/Borel.thy Probability/Measure.thy \
+ Probability/Lebesgue.thy Probability/Product_Measure.thy \
Probability/Probability_Space.thy
- @$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL Probability
+ @cd Probability; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL HOL-Probability
## HOL-Nominal
@@ -1233,7 +1232,7 @@
SMT/Tools/yices_solver.ML SMT/Tools/z3_proof_terms.ML \
SMT/Tools/z3_proof_rules.ML SMT/Tools/z3_proof.ML \
SMT/Tools/z3_model.ML SMT/Tools/z3_interface.ML \
- SMT/Tools/z3_solver.ML $(SRC)/Tools/Cache_IO/cache_io.ML
+ SMT/Tools/z3_solver.ML $(SRC)/Tools/cache_io.ML
@cd SMT; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL-Word HOL-SMT
@@ -1274,7 +1273,7 @@
HOL-Mutabelle: HOL $(LOG)/HOL-Mutabelle.gz
$(LOG)/HOL-Mutabelle.gz: $(OUT)/HOL Mutabelle/MutabelleExtra.thy \
- Mutabelle/ROOT.ML Mutabelle/mutabelle.ML \
+ Mutabelle/ROOT.ML Mutabelle/mutabelle.ML \
Mutabelle/mutabelle_extra.ML
@$(ISABELLE_TOOL) usedir $(OUT)/HOL Mutabelle
@@ -1283,11 +1282,22 @@
HOL-Quotient_Examples: HOL $(LOG)/HOL-Quotient_Examples.gz
-$(LOG)/HOL-Quotient_Examples.gz: $(OUT)/HOL \
+$(LOG)/HOL-Quotient_Examples.gz: $(OUT)/HOL \
Quotient_Examples/LarryInt.thy Quotient_Examples/LarryDatatype.thy
@$(ISABELLE_TOOL) usedir $(OUT)/HOL Quotient_Examples
+## HOL-Predicate_Compile_Examples
+
+HOL-Predicate_Compile_Examples: HOL $(LOG)/HOL-Predicate_Compile_Examples.gz
+
+$(LOG)/HOL-Predicate_Compile_Examples.gz: $(OUT)/HOL \
+ Predicate_Compile_Examples/ROOT.ML \
+ Predicate_Compile_Examples/Predicate_Compile_Examples.thy \
+ Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy
+ @$(ISABELLE_TOOL) usedir $(OUT)/HOL Predicate_Compile_Examples
+
+
## clean
clean:
@@ -1311,18 +1321,19 @@
$(LOG)/HOL-Nominal-Examples.gz $(LOG)/HOL-Nominal.gz \
$(LOG)/HOL-Number_Theory.gz \
$(LOG)/HOL-Old_Number_Theory.gz $(LOG)/HOL-Plain.gz \
+ $(LOG)/HOL-Predicate_Compile_Examples.gz \
$(LOG)/HOL-Probability.gz $(LOG)/HOL-Prolog.gz \
$(LOG)/HOL-Proofs.gz $(LOG)/HOL-Proofs-Extraction.gz \
$(LOG)/HOL-Proofs-Lambda.gz $(LOG)/HOL-SET_Protocol.gz \
$(LOG)/HOL-SMT-Examples.gz $(LOG)/HOL-SMT.gz \
$(LOG)/HOL-Statespace.gz $(LOG)/HOL-Subst.gz \
$(LOG)/HOL-UNITY.gz $(LOG)/HOL-Unix.gz \
- $(LOG)/HOL-Word-Examples.gz \
- $(LOG)/HOL-Word.gz $(LOG)/HOL-ZF.gz $(LOG)/HOL-ex.gz \
- $(LOG)/HOL.gz $(LOG)/HOL4.gz $(LOG)/TLA-Buffer.gz \
- $(LOG)/TLA-Inc.gz $(LOG)/TLA-Memory.gz $(LOG)/TLA.gz \
- $(OUT)/HOL $(OUT)/HOL-Algebra $(OUT)/HOL-Base \
- $(OUT)/HOL-Boogie $(OUT)/HOL-Main \
- $(OUT)/HOL-Multivariate_Analysis $(OUT)/HOL-NSA \
- $(OUT)/HOL-Nominal $(OUT)/HOL-Plain $(OUT)/HOL-Proofs \
+ $(LOG)/HOL-Word-Examples.gz $(LOG)/HOL-Word.gz \
+ $(LOG)/HOL-ZF.gz $(LOG)/HOL-ex.gz $(LOG)/HOL.gz \
+ $(LOG)/HOL4.gz $(LOG)/TLA-Buffer.gz $(LOG)/TLA-Inc.gz \
+ $(LOG)/TLA-Memory.gz $(LOG)/TLA.gz $(OUT)/HOL \
+ $(OUT)/HOL-Algebra $(OUT)/HOL-Base $(OUT)/HOL-Boogie \
+ $(OUT)/HOL-Main $(OUT)/HOL-Multivariate_Analysis \
+ $(OUT)/HOL-NSA $(OUT)/HOL-Nominal $(OUT)/HOL-Plain \
+ $(OUT)/HOL-Probability $(OUT)/HOL-Proofs \
$(OUT)/HOL-SMT $(OUT)/HOL-Word $(OUT)/HOL4 $(OUT)/TLA
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy Thu Mar 25 17:56:31 2010 +0100
@@ -0,0 +1,147 @@
+theory Predicate_Compile_Alternative_Defs
+imports "../Predicate_Compile"
+begin
+
+section {* Common constants *}
+
+declare HOL.if_bool_eq_disj[code_pred_inline]
+
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name Let}] *}
+
+section {* Pairs *}
+
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name fst}, @{const_name snd}, @{const_name split}] *}
+
+section {* Bounded quantifiers *}
+
+declare Ball_def[code_pred_inline]
+declare Bex_def[code_pred_inline]
+
+section {* Set operations *}
+
+declare Collect_def[code_pred_inline]
+declare mem_def[code_pred_inline]
+
+declare eq_reflection[OF empty_def, code_pred_inline]
+declare insert_code[code_pred_def]
+
+declare subset_iff[code_pred_inline]
+
+declare Int_def[code_pred_inline]
+declare eq_reflection[OF Un_def, code_pred_inline]
+declare eq_reflection[OF UNION_def, code_pred_inline]
+
+lemma Diff[code_pred_inline]:
+ "(A - B) = (%x. A x \<and> \<not> B x)"
+by (auto simp add: mem_def)
+
+lemma set_equality[code_pred_inline]:
+ "(A = B) = ((\<forall>x. A x \<longrightarrow> B x) \<and> (\<forall>x. B x \<longrightarrow> A x))"
+by (fastsimp simp add: mem_def)
+
+section {* Setup for Numerals *}
+
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name number_of}] *}
+setup {* Predicate_Compile_Data.keep_functions [@{const_name number_of}] *}
+
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name div}, @{const_name mod}, @{const_name times}] *}
+
+subsection {* Inductive definitions for arithmetic on natural numbers *}
+
+inductive plusP
+where
+ "plusP x 0 x"
+| "plusP x y z ==> plusP x (Suc y) (Suc z)"
+
+setup {* Predicate_Compile_Fun.add_function_predicate_translation
+ (@{term "op + :: nat => nat => nat"}, @{term "plusP"}) *}
+
+inductive less_nat
+where
+ "less_nat 0 (Suc y)"
+| "less_nat x y ==> less_nat (Suc x) (Suc y)"
+
+lemma [code_pred_inline]:
+ "x < y = less_nat x y"
+apply (rule iffI)
+apply (induct x arbitrary: y)
+apply (case_tac y) apply (auto intro: less_nat.intros)
+apply (case_tac y)
+apply (auto intro: less_nat.intros)
+apply (induct rule: less_nat.induct)
+apply auto
+done
+
+inductive less_eq_nat
+where
+ "less_eq_nat 0 y"
+| "less_eq_nat x y ==> less_eq_nat (Suc x) (Suc y)"
+
+lemma [code_pred_inline]:
+"x <= y = less_eq_nat x y"
+apply (rule iffI)
+apply (induct x arbitrary: y)
+apply (auto intro: less_eq_nat.intros)
+apply (case_tac y) apply (auto intro: less_eq_nat.intros)
+apply (induct rule: less_eq_nat.induct)
+apply auto done
+
+section {* Alternative list definitions *}
+
+text {* size simps are not yet added to the Spec_Rules interface. So they are just added manually here! *}
+
+lemma [code_pred_def]:
+ "length [] = 0"
+ "length (x # xs) = Suc (length xs)"
+by auto
+
+subsection {* Alternative rules for set *}
+
+lemma set_ConsI1 [code_pred_intro]:
+ "set (x # xs) x"
+unfolding mem_def[symmetric, of _ x]
+by auto
+
+lemma set_ConsI2 [code_pred_intro]:
+ "set xs x ==> set (x' # xs) x"
+unfolding mem_def[symmetric, of _ x]
+by auto
+
+code_pred [skip_proof] set
+proof -
+ case set
+ from this show thesis
+ apply (case_tac xb)
+ apply auto
+ unfolding mem_def[symmetric, of _ xc]
+ apply auto
+ unfolding mem_def
+ apply fastsimp
+ done
+qed
+
+subsection {* Alternative rules for list_all2 *}
+
+lemma list_all2_NilI [code_pred_intro]: "list_all2 P [] []"
+by auto
+
+lemma list_all2_ConsI [code_pred_intro]: "list_all2 P xs ys ==> P x y ==> list_all2 P (x#xs) (y#ys)"
+by auto
+
+code_pred [skip_proof] list_all2
+proof -
+ case list_all2
+ from this show thesis
+ apply -
+ apply (case_tac xb)
+ apply (case_tac xc)
+ apply auto
+ apply (case_tac xc)
+ apply auto
+ apply fastsimp
+ done
+qed
+
+
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Predicate_Compile_Quickcheck.thy Thu Mar 25 17:56:31 2010 +0100
@@ -0,0 +1,14 @@
+(* Author: Lukas Bulwahn, TU Muenchen *)
+
+header {* A Prototype of Quickcheck based on the Predicate Compiler *}
+
+theory Predicate_Compile_Quickcheck
+imports Main Predicate_Compile_Alternative_Defs
+uses "../Tools/Predicate_Compile/predicate_compile_quickcheck.ML"
+begin
+
+setup {* Quickcheck.add_generator ("predicate_compile_wo_ff", Predicate_Compile_Quickcheck.quickcheck_compile_term false true 4) *}
+setup {* Quickcheck.add_generator ("predicate_compile_ff_fs", Predicate_Compile_Quickcheck.quickcheck_compile_term true true 4) *}
+setup {* Quickcheck.add_generator ("predicate_compile_ff_nofs", Predicate_Compile_Quickcheck.quickcheck_compile_term true false 4) *}
+
+end
\ No newline at end of file
--- a/src/HOL/Multivariate_Analysis/Integration.cert Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOL/Multivariate_Analysis/Integration.cert Thu Mar 25 17:56:31 2010 +0100
@@ -1,4 +1,853 @@
-tB2Atlor9W4pSnrAz5nHpw 907 0
+534d3afbd3505abffa6b03e3986bfa7aa3cff358 428 0
+#2 := false
+decl uf_10 :: T1
+#38 := uf_10
+decl uf_3 :: T1
+#21 := uf_3
+#45 := (= uf_3 uf_10)
+decl uf_1 :: (-> int T1)
+decl uf_2 :: (-> T1 int)
+#39 := (uf_2 uf_10)
+#588 := (uf_1 #39)
+#686 := (= #588 uf_10)
+#589 := (= uf_10 #588)
+#4 := (:var 0 T1)
+#5 := (uf_2 #4)
+#541 := (pattern #5)
+#6 := (uf_1 #5)
+#93 := (= #4 #6)
+#542 := (forall (vars (?x1 T1)) (:pat #541) #93)
+#96 := (forall (vars (?x1 T1)) #93)
+#545 := (iff #96 #542)
+#543 := (iff #93 #93)
+#544 := [refl]: #543
+#546 := [quant-intro #544]: #545
+#454 := (~ #96 #96)
+#456 := (~ #93 #93)
+#457 := [refl]: #456
+#455 := [nnf-pos #457]: #454
+#7 := (= #6 #4)
+#8 := (forall (vars (?x1 T1)) #7)
+#97 := (iff #8 #96)
+#94 := (iff #7 #93)
+#95 := [rewrite]: #94
+#98 := [quant-intro #95]: #97
+#92 := [asserted]: #8
+#101 := [mp #92 #98]: #96
+#452 := [mp~ #101 #455]: #96
+#547 := [mp #452 #546]: #542
+#590 := (not #542)
+#595 := (or #590 #589)
+#596 := [quant-inst]: #595
+#680 := [unit-resolution #596 #547]: #589
+#687 := [symm #680]: #686
+#688 := (= uf_3 #588)
+#22 := (uf_2 uf_3)
+#586 := (uf_1 #22)
+#684 := (= #586 #588)
+#682 := (= #588 #586)
+#678 := (= #39 #22)
+#676 := (= #22 #39)
+#9 := 0::int
+#227 := -1::int
+#230 := (* -1::int #39)
+#231 := (+ #22 #230)
+#296 := (<= #231 0::int)
+#70 := (<= #22 #39)
+#393 := (iff #70 #296)
+#394 := [rewrite]: #393
+#347 := [asserted]: #70
+#395 := [mp #347 #394]: #296
+#229 := (>= #231 0::int)
+decl uf_4 :: (-> T2 T3 real)
+decl uf_6 :: (-> T1 T3)
+#25 := (uf_6 uf_3)
+decl uf_7 :: T2
+#27 := uf_7
+#28 := (uf_4 uf_7 #25)
+decl uf_9 :: T2
+#33 := uf_9
+#34 := (uf_4 uf_9 #25)
+#46 := (uf_6 uf_10)
+decl uf_5 :: T2
+#24 := uf_5
+#47 := (uf_4 uf_5 #46)
+#48 := (ite #45 #47 #34)
+#256 := (ite #229 #48 #28)
+#568 := (= #28 #256)
+#648 := (not #568)
+#194 := 0::real
+#192 := -1::real
+#265 := (* -1::real #256)
+#640 := (+ #28 #265)
+#642 := (>= #640 0::real)
+#645 := (not #642)
+#643 := [hypothesis]: #642
+decl uf_8 :: T2
+#30 := uf_8
+#31 := (uf_4 uf_8 #25)
+#266 := (+ #31 #265)
+#264 := (>= #266 0::real)
+#267 := (not #264)
+#26 := (uf_4 uf_5 #25)
+decl uf_11 :: T2
+#41 := uf_11
+#42 := (uf_4 uf_11 #25)
+#237 := (ite #229 #42 #26)
+#245 := (* -1::real #237)
+#246 := (+ #31 #245)
+#247 := (<= #246 0::real)
+#248 := (not #247)
+#272 := (and #248 #267)
+#40 := (< #22 #39)
+#49 := (ite #40 #28 #48)
+#50 := (< #31 #49)
+#43 := (ite #40 #26 #42)
+#44 := (< #43 #31)
+#51 := (and #44 #50)
+#273 := (iff #51 #272)
+#270 := (iff #50 #267)
+#261 := (< #31 #256)
+#268 := (iff #261 #267)
+#269 := [rewrite]: #268
+#262 := (iff #50 #261)
+#259 := (= #49 #256)
+#228 := (not #229)
+#253 := (ite #228 #28 #48)
+#257 := (= #253 #256)
+#258 := [rewrite]: #257
+#254 := (= #49 #253)
+#232 := (iff #40 #228)
+#233 := [rewrite]: #232
+#255 := [monotonicity #233]: #254
+#260 := [trans #255 #258]: #259
+#263 := [monotonicity #260]: #262
+#271 := [trans #263 #269]: #270
+#251 := (iff #44 #248)
+#242 := (< #237 #31)
+#249 := (iff #242 #248)
+#250 := [rewrite]: #249
+#243 := (iff #44 #242)
+#240 := (= #43 #237)
+#234 := (ite #228 #26 #42)
+#238 := (= #234 #237)
+#239 := [rewrite]: #238
+#235 := (= #43 #234)
+#236 := [monotonicity #233]: #235
+#241 := [trans #236 #239]: #240
+#244 := [monotonicity #241]: #243
+#252 := [trans #244 #250]: #251
+#274 := [monotonicity #252 #271]: #273
+#178 := [asserted]: #51
+#275 := [mp #178 #274]: #272
+#277 := [and-elim #275]: #267
+#196 := (* -1::real #31)
+#197 := (+ #28 #196)
+#195 := (>= #197 0::real)
+#193 := (not #195)
+#213 := (* -1::real #34)
+#214 := (+ #31 #213)
+#212 := (>= #214 0::real)
+#215 := (not #212)
+#220 := (and #193 #215)
+#23 := (< #22 #22)
+#35 := (ite #23 #28 #34)
+#36 := (< #31 #35)
+#29 := (ite #23 #26 #28)
+#32 := (< #29 #31)
+#37 := (and #32 #36)
+#221 := (iff #37 #220)
+#218 := (iff #36 #215)
+#209 := (< #31 #34)
+#216 := (iff #209 #215)
+#217 := [rewrite]: #216
+#210 := (iff #36 #209)
+#207 := (= #35 #34)
+#202 := (ite false #28 #34)
+#205 := (= #202 #34)
+#206 := [rewrite]: #205
+#203 := (= #35 #202)
+#180 := (iff #23 false)
+#181 := [rewrite]: #180
+#204 := [monotonicity #181]: #203
+#208 := [trans #204 #206]: #207
+#211 := [monotonicity #208]: #210
+#219 := [trans #211 #217]: #218
+#200 := (iff #32 #193)
+#189 := (< #28 #31)
+#198 := (iff #189 #193)
+#199 := [rewrite]: #198
+#190 := (iff #32 #189)
+#187 := (= #29 #28)
+#182 := (ite false #26 #28)
+#185 := (= #182 #28)
+#186 := [rewrite]: #185
+#183 := (= #29 #182)
+#184 := [monotonicity #181]: #183
+#188 := [trans #184 #186]: #187
+#191 := [monotonicity #188]: #190
+#201 := [trans #191 #199]: #200
+#222 := [monotonicity #201 #219]: #221
+#177 := [asserted]: #37
+#223 := [mp #177 #222]: #220
+#224 := [and-elim #223]: #193
+#644 := [th-lemma #224 #277 #643]: false
+#646 := [lemma #644]: #645
+#647 := [hypothesis]: #568
+#649 := (or #648 #642)
+#650 := [th-lemma]: #649
+#651 := [unit-resolution #650 #647 #646]: false
+#652 := [lemma #651]: #648
+#578 := (or #229 #568)
+#579 := [def-axiom]: #578
+#675 := [unit-resolution #579 #652]: #229
+#677 := [th-lemma #675 #395]: #676
+#679 := [symm #677]: #678
+#683 := [monotonicity #679]: #682
+#685 := [symm #683]: #684
+#587 := (= uf_3 #586)
+#591 := (or #590 #587)
+#592 := [quant-inst]: #591
+#681 := [unit-resolution #592 #547]: #587
+#689 := [trans #681 #685]: #688
+#690 := [trans #689 #687]: #45
+#571 := (not #45)
+#54 := (uf_4 uf_11 #46)
+#279 := (ite #45 #28 #54)
+#465 := (* -1::real #279)
+#632 := (+ #28 #465)
+#633 := (<= #632 0::real)
+#580 := (= #28 #279)
+#656 := [hypothesis]: #45
+#582 := (or #571 #580)
+#583 := [def-axiom]: #582
+#657 := [unit-resolution #583 #656]: #580
+#658 := (not #580)
+#659 := (or #658 #633)
+#660 := [th-lemma]: #659
+#661 := [unit-resolution #660 #657]: #633
+#57 := (uf_4 uf_8 #46)
+#363 := (* -1::real #57)
+#379 := (+ #47 #363)
+#380 := (<= #379 0::real)
+#381 := (not #380)
+#364 := (+ #54 #363)
+#362 := (>= #364 0::real)
+#361 := (not #362)
+#386 := (and #361 #381)
+#59 := (uf_4 uf_7 #46)
+#64 := (< #39 #39)
+#67 := (ite #64 #59 #47)
+#68 := (< #57 #67)
+#65 := (ite #64 #47 #54)
+#66 := (< #65 #57)
+#69 := (and #66 #68)
+#387 := (iff #69 #386)
+#384 := (iff #68 #381)
+#376 := (< #57 #47)
+#382 := (iff #376 #381)
+#383 := [rewrite]: #382
+#377 := (iff #68 #376)
+#374 := (= #67 #47)
+#369 := (ite false #59 #47)
+#372 := (= #369 #47)
+#373 := [rewrite]: #372
+#370 := (= #67 #369)
+#349 := (iff #64 false)
+#350 := [rewrite]: #349
+#371 := [monotonicity #350]: #370
+#375 := [trans #371 #373]: #374
+#378 := [monotonicity #375]: #377
+#385 := [trans #378 #383]: #384
+#367 := (iff #66 #361)
+#358 := (< #54 #57)
+#365 := (iff #358 #361)
+#366 := [rewrite]: #365
+#359 := (iff #66 #358)
+#356 := (= #65 #54)
+#351 := (ite false #47 #54)
+#354 := (= #351 #54)
+#355 := [rewrite]: #354
+#352 := (= #65 #351)
+#353 := [monotonicity #350]: #352
+#357 := [trans #353 #355]: #356
+#360 := [monotonicity #357]: #359
+#368 := [trans #360 #366]: #367
+#388 := [monotonicity #368 #385]: #387
+#346 := [asserted]: #69
+#389 := [mp #346 #388]: #386
+#391 := [and-elim #389]: #381
+#397 := (* -1::real #59)
+#398 := (+ #47 #397)
+#399 := (<= #398 0::real)
+#409 := (* -1::real #54)
+#410 := (+ #47 #409)
+#408 := (>= #410 0::real)
+#60 := (uf_4 uf_9 #46)
+#402 := (* -1::real #60)
+#403 := (+ #59 #402)
+#404 := (<= #403 0::real)
+#418 := (and #399 #404 #408)
+#73 := (<= #59 #60)
+#72 := (<= #47 #59)
+#74 := (and #72 #73)
+#71 := (<= #54 #47)
+#75 := (and #71 #74)
+#421 := (iff #75 #418)
+#412 := (and #399 #404)
+#415 := (and #408 #412)
+#419 := (iff #415 #418)
+#420 := [rewrite]: #419
+#416 := (iff #75 #415)
+#413 := (iff #74 #412)
+#405 := (iff #73 #404)
+#406 := [rewrite]: #405
+#400 := (iff #72 #399)
+#401 := [rewrite]: #400
+#414 := [monotonicity #401 #406]: #413
+#407 := (iff #71 #408)
+#411 := [rewrite]: #407
+#417 := [monotonicity #411 #414]: #416
+#422 := [trans #417 #420]: #421
+#348 := [asserted]: #75
+#423 := [mp #348 #422]: #418
+#424 := [and-elim #423]: #399
+#637 := (+ #28 #397)
+#639 := (>= #637 0::real)
+#636 := (= #28 #59)
+#666 := (= #59 #28)
+#664 := (= #46 #25)
+#662 := (= #25 #46)
+#663 := [monotonicity #656]: #662
+#665 := [symm #663]: #664
+#667 := [monotonicity #665]: #666
+#668 := [symm #667]: #636
+#669 := (not #636)
+#670 := (or #669 #639)
+#671 := [th-lemma]: #670
+#672 := [unit-resolution #671 #668]: #639
+#468 := (+ #57 #465)
+#471 := (<= #468 0::real)
+#444 := (not #471)
+#322 := (ite #296 #279 #47)
+#330 := (* -1::real #322)
+#331 := (+ #57 #330)
+#332 := (<= #331 0::real)
+#333 := (not #332)
+#445 := (iff #333 #444)
+#472 := (iff #332 #471)
+#469 := (= #331 #468)
+#466 := (= #330 #465)
+#463 := (= #322 #279)
+#1 := true
+#458 := (ite true #279 #47)
+#461 := (= #458 #279)
+#462 := [rewrite]: #461
+#459 := (= #322 #458)
+#450 := (iff #296 true)
+#451 := [iff-true #395]: #450
+#460 := [monotonicity #451]: #459
+#464 := [trans #460 #462]: #463
+#467 := [monotonicity #464]: #466
+#470 := [monotonicity #467]: #469
+#473 := [monotonicity #470]: #472
+#474 := [monotonicity #473]: #445
+#303 := (ite #296 #60 #59)
+#313 := (* -1::real #303)
+#314 := (+ #57 #313)
+#312 := (>= #314 0::real)
+#311 := (not #312)
+#338 := (and #311 #333)
+#52 := (< #39 #22)
+#61 := (ite #52 #59 #60)
+#62 := (< #57 #61)
+#53 := (= uf_10 uf_3)
+#55 := (ite #53 #28 #54)
+#56 := (ite #52 #47 #55)
+#58 := (< #56 #57)
+#63 := (and #58 #62)
+#341 := (iff #63 #338)
+#282 := (ite #52 #47 #279)
+#285 := (< #282 #57)
+#291 := (and #62 #285)
+#339 := (iff #291 #338)
+#336 := (iff #285 #333)
+#327 := (< #322 #57)
+#334 := (iff #327 #333)
+#335 := [rewrite]: #334
+#328 := (iff #285 #327)
+#325 := (= #282 #322)
+#297 := (not #296)
+#319 := (ite #297 #47 #279)
+#323 := (= #319 #322)
+#324 := [rewrite]: #323
+#320 := (= #282 #319)
+#298 := (iff #52 #297)
+#299 := [rewrite]: #298
+#321 := [monotonicity #299]: #320
+#326 := [trans #321 #324]: #325
+#329 := [monotonicity #326]: #328
+#337 := [trans #329 #335]: #336
+#317 := (iff #62 #311)
+#308 := (< #57 #303)
+#315 := (iff #308 #311)
+#316 := [rewrite]: #315
+#309 := (iff #62 #308)
+#306 := (= #61 #303)
+#300 := (ite #297 #59 #60)
+#304 := (= #300 #303)
+#305 := [rewrite]: #304
+#301 := (= #61 #300)
+#302 := [monotonicity #299]: #301
+#307 := [trans #302 #305]: #306
+#310 := [monotonicity #307]: #309
+#318 := [trans #310 #316]: #317
+#340 := [monotonicity #318 #337]: #339
+#294 := (iff #63 #291)
+#288 := (and #285 #62)
+#292 := (iff #288 #291)
+#293 := [rewrite]: #292
+#289 := (iff #63 #288)
+#286 := (iff #58 #285)
+#283 := (= #56 #282)
+#280 := (= #55 #279)
+#226 := (iff #53 #45)
+#278 := [rewrite]: #226
+#281 := [monotonicity #278]: #280
+#284 := [monotonicity #281]: #283
+#287 := [monotonicity #284]: #286
+#290 := [monotonicity #287]: #289
+#295 := [trans #290 #293]: #294
+#342 := [trans #295 #340]: #341
+#179 := [asserted]: #63
+#343 := [mp #179 #342]: #338
+#345 := [and-elim #343]: #333
+#475 := [mp #345 #474]: #444
+#673 := [th-lemma #475 #672 #424 #391 #661]: false
+#674 := [lemma #673]: #571
+[unit-resolution #674 #690]: false
+unsat
+985f86fd1ea01264c02b8076ffd658677bd4d29a 419 0
+#2 := false
+#194 := 0::real
+decl uf_4 :: (-> T2 T3 real)
+decl uf_6 :: (-> T1 T3)
+decl uf_3 :: T1
+#21 := uf_3
+#25 := (uf_6 uf_3)
+decl uf_5 :: T2
+#24 := uf_5
+#26 := (uf_4 uf_5 #25)
+decl uf_7 :: T2
+#27 := uf_7
+#28 := (uf_4 uf_7 #25)
+decl uf_10 :: T1
+#38 := uf_10
+#42 := (uf_6 uf_10)
+decl uf_9 :: T2
+#33 := uf_9
+#43 := (uf_4 uf_9 #42)
+#41 := (= uf_3 uf_10)
+#44 := (ite #41 #43 #28)
+#9 := 0::int
+decl uf_2 :: (-> T1 int)
+#39 := (uf_2 uf_10)
+#226 := -1::int
+#229 := (* -1::int #39)
+#22 := (uf_2 uf_3)
+#230 := (+ #22 #229)
+#228 := (>= #230 0::int)
+#236 := (ite #228 #44 #26)
+#192 := -1::real
+#244 := (* -1::real #236)
+#642 := (+ #26 #244)
+#643 := (<= #642 0::real)
+#567 := (= #26 #236)
+#227 := (not #228)
+decl uf_1 :: (-> int T1)
+#593 := (uf_1 #39)
+#660 := (= #593 uf_10)
+#594 := (= uf_10 #593)
+#4 := (:var 0 T1)
+#5 := (uf_2 #4)
+#546 := (pattern #5)
+#6 := (uf_1 #5)
+#93 := (= #4 #6)
+#547 := (forall (vars (?x1 T1)) (:pat #546) #93)
+#96 := (forall (vars (?x1 T1)) #93)
+#550 := (iff #96 #547)
+#548 := (iff #93 #93)
+#549 := [refl]: #548
+#551 := [quant-intro #549]: #550
+#448 := (~ #96 #96)
+#450 := (~ #93 #93)
+#451 := [refl]: #450
+#449 := [nnf-pos #451]: #448
+#7 := (= #6 #4)
+#8 := (forall (vars (?x1 T1)) #7)
+#97 := (iff #8 #96)
+#94 := (iff #7 #93)
+#95 := [rewrite]: #94
+#98 := [quant-intro #95]: #97
+#92 := [asserted]: #8
+#101 := [mp #92 #98]: #96
+#446 := [mp~ #101 #449]: #96
+#552 := [mp #446 #551]: #547
+#595 := (not #547)
+#600 := (or #595 #594)
+#601 := [quant-inst]: #600
+#654 := [unit-resolution #601 #552]: #594
+#680 := [symm #654]: #660
+#681 := (= uf_3 #593)
+#591 := (uf_1 #22)
+#658 := (= #591 #593)
+#656 := (= #593 #591)
+#652 := (= #39 #22)
+#647 := (= #22 #39)
+#290 := (<= #230 0::int)
+#70 := (<= #22 #39)
+#388 := (iff #70 #290)
+#389 := [rewrite]: #388
+#341 := [asserted]: #70
+#390 := [mp #341 #389]: #290
+#646 := [hypothesis]: #228
+#648 := [th-lemma #646 #390]: #647
+#653 := [symm #648]: #652
+#657 := [monotonicity #653]: #656
+#659 := [symm #657]: #658
+#592 := (= uf_3 #591)
+#596 := (or #595 #592)
+#597 := [quant-inst]: #596
+#655 := [unit-resolution #597 #552]: #592
+#682 := [trans #655 #659]: #681
+#683 := [trans #682 #680]: #41
+#570 := (not #41)
+decl uf_11 :: T2
+#47 := uf_11
+#59 := (uf_4 uf_11 #42)
+#278 := (ite #41 #26 #59)
+#459 := (* -1::real #278)
+#637 := (+ #26 #459)
+#639 := (>= #637 0::real)
+#585 := (= #26 #278)
+#661 := [hypothesis]: #41
+#587 := (or #570 #585)
+#588 := [def-axiom]: #587
+#662 := [unit-resolution #588 #661]: #585
+#663 := (not #585)
+#664 := (or #663 #639)
+#665 := [th-lemma]: #664
+#666 := [unit-resolution #665 #662]: #639
+decl uf_8 :: T2
+#30 := uf_8
+#56 := (uf_4 uf_8 #42)
+#357 := (* -1::real #56)
+#358 := (+ #43 #357)
+#356 := (>= #358 0::real)
+#355 := (not #356)
+#374 := (* -1::real #59)
+#375 := (+ #56 #374)
+#373 := (>= #375 0::real)
+#376 := (not #373)
+#381 := (and #355 #376)
+#64 := (< #39 #39)
+#67 := (ite #64 #43 #59)
+#68 := (< #56 #67)
+#53 := (uf_4 uf_5 #42)
+#65 := (ite #64 #53 #43)
+#66 := (< #65 #56)
+#69 := (and #66 #68)
+#382 := (iff #69 #381)
+#379 := (iff #68 #376)
+#370 := (< #56 #59)
+#377 := (iff #370 #376)
+#378 := [rewrite]: #377
+#371 := (iff #68 #370)
+#368 := (= #67 #59)
+#363 := (ite false #43 #59)
+#366 := (= #363 #59)
+#367 := [rewrite]: #366
+#364 := (= #67 #363)
+#343 := (iff #64 false)
+#344 := [rewrite]: #343
+#365 := [monotonicity #344]: #364
+#369 := [trans #365 #367]: #368
+#372 := [monotonicity #369]: #371
+#380 := [trans #372 #378]: #379
+#361 := (iff #66 #355)
+#352 := (< #43 #56)
+#359 := (iff #352 #355)
+#360 := [rewrite]: #359
+#353 := (iff #66 #352)
+#350 := (= #65 #43)
+#345 := (ite false #53 #43)
+#348 := (= #345 #43)
+#349 := [rewrite]: #348
+#346 := (= #65 #345)
+#347 := [monotonicity #344]: #346
+#351 := [trans #347 #349]: #350
+#354 := [monotonicity #351]: #353
+#362 := [trans #354 #360]: #361
+#383 := [monotonicity #362 #380]: #382
+#340 := [asserted]: #69
+#384 := [mp #340 #383]: #381
+#385 := [and-elim #384]: #355
+#394 := (* -1::real #53)
+#395 := (+ #43 #394)
+#393 := (>= #395 0::real)
+#54 := (uf_4 uf_7 #42)
+#402 := (* -1::real #54)
+#403 := (+ #53 #402)
+#401 := (>= #403 0::real)
+#397 := (+ #43 #374)
+#398 := (<= #397 0::real)
+#412 := (and #393 #398 #401)
+#73 := (<= #43 #59)
+#72 := (<= #53 #43)
+#74 := (and #72 #73)
+#71 := (<= #54 #53)
+#75 := (and #71 #74)
+#415 := (iff #75 #412)
+#406 := (and #393 #398)
+#409 := (and #401 #406)
+#413 := (iff #409 #412)
+#414 := [rewrite]: #413
+#410 := (iff #75 #409)
+#407 := (iff #74 #406)
+#399 := (iff #73 #398)
+#400 := [rewrite]: #399
+#392 := (iff #72 #393)
+#396 := [rewrite]: #392
+#408 := [monotonicity #396 #400]: #407
+#404 := (iff #71 #401)
+#405 := [rewrite]: #404
+#411 := [monotonicity #405 #408]: #410
+#416 := [trans #411 #414]: #415
+#342 := [asserted]: #75
+#417 := [mp #342 #416]: #412
+#418 := [and-elim #417]: #393
+#650 := (+ #26 #394)
+#651 := (<= #650 0::real)
+#649 := (= #26 #53)
+#671 := (= #53 #26)
+#669 := (= #42 #25)
+#667 := (= #25 #42)
+#668 := [monotonicity #661]: #667
+#670 := [symm #668]: #669
+#672 := [monotonicity #670]: #671
+#673 := [symm #672]: #649
+#674 := (not #649)
+#675 := (or #674 #651)
+#676 := [th-lemma]: #675
+#677 := [unit-resolution #676 #673]: #651
+#462 := (+ #56 #459)
+#465 := (>= #462 0::real)
+#438 := (not #465)
+#316 := (ite #290 #278 #43)
+#326 := (* -1::real #316)
+#327 := (+ #56 #326)
+#325 := (>= #327 0::real)
+#324 := (not #325)
+#439 := (iff #324 #438)
+#466 := (iff #325 #465)
+#463 := (= #327 #462)
+#460 := (= #326 #459)
+#457 := (= #316 #278)
+#1 := true
+#452 := (ite true #278 #43)
+#455 := (= #452 #278)
+#456 := [rewrite]: #455
+#453 := (= #316 #452)
+#444 := (iff #290 true)
+#445 := [iff-true #390]: #444
+#454 := [monotonicity #445]: #453
+#458 := [trans #454 #456]: #457
+#461 := [monotonicity #458]: #460
+#464 := [monotonicity #461]: #463
+#467 := [monotonicity #464]: #466
+#468 := [monotonicity #467]: #439
+#297 := (ite #290 #54 #53)
+#305 := (* -1::real #297)
+#306 := (+ #56 #305)
+#307 := (<= #306 0::real)
+#308 := (not #307)
+#332 := (and #308 #324)
+#58 := (= uf_10 uf_3)
+#60 := (ite #58 #26 #59)
+#52 := (< #39 #22)
+#61 := (ite #52 #43 #60)
+#62 := (< #56 #61)
+#55 := (ite #52 #53 #54)
+#57 := (< #55 #56)
+#63 := (and #57 #62)
+#335 := (iff #63 #332)
+#281 := (ite #52 #43 #278)
+#284 := (< #56 #281)
+#287 := (and #57 #284)
+#333 := (iff #287 #332)
+#330 := (iff #284 #324)
+#321 := (< #56 #316)
+#328 := (iff #321 #324)
+#329 := [rewrite]: #328
+#322 := (iff #284 #321)
+#319 := (= #281 #316)
+#291 := (not #290)
+#313 := (ite #291 #43 #278)
+#317 := (= #313 #316)
+#318 := [rewrite]: #317
+#314 := (= #281 #313)
+#292 := (iff #52 #291)
+#293 := [rewrite]: #292
+#315 := [monotonicity #293]: #314
+#320 := [trans #315 #318]: #319
+#323 := [monotonicity #320]: #322
+#331 := [trans #323 #329]: #330
+#311 := (iff #57 #308)
+#302 := (< #297 #56)
+#309 := (iff #302 #308)
+#310 := [rewrite]: #309
+#303 := (iff #57 #302)
+#300 := (= #55 #297)
+#294 := (ite #291 #53 #54)
+#298 := (= #294 #297)
+#299 := [rewrite]: #298
+#295 := (= #55 #294)
+#296 := [monotonicity #293]: #295
+#301 := [trans #296 #299]: #300
+#304 := [monotonicity #301]: #303
+#312 := [trans #304 #310]: #311
+#334 := [monotonicity #312 #331]: #333
+#288 := (iff #63 #287)
+#285 := (iff #62 #284)
+#282 := (= #61 #281)
+#279 := (= #60 #278)
+#225 := (iff #58 #41)
+#277 := [rewrite]: #225
+#280 := [monotonicity #277]: #279
+#283 := [monotonicity #280]: #282
+#286 := [monotonicity #283]: #285
+#289 := [monotonicity #286]: #288
+#336 := [trans #289 #334]: #335
+#179 := [asserted]: #63
+#337 := [mp #179 #336]: #332
+#339 := [and-elim #337]: #324
+#469 := [mp #339 #468]: #438
+#678 := [th-lemma #469 #677 #418 #385 #666]: false
+#679 := [lemma #678]: #570
+#684 := [unit-resolution #679 #683]: false
+#685 := [lemma #684]: #227
+#577 := (or #228 #567)
+#578 := [def-axiom]: #577
+#645 := [unit-resolution #578 #685]: #567
+#686 := (not #567)
+#687 := (or #686 #643)
+#688 := [th-lemma]: #687
+#689 := [unit-resolution #688 #645]: #643
+#31 := (uf_4 uf_8 #25)
+#245 := (+ #31 #244)
+#246 := (<= #245 0::real)
+#247 := (not #246)
+#34 := (uf_4 uf_9 #25)
+#48 := (uf_4 uf_11 #25)
+#255 := (ite #228 #48 #34)
+#264 := (* -1::real #255)
+#265 := (+ #31 #264)
+#263 := (>= #265 0::real)
+#266 := (not #263)
+#271 := (and #247 #266)
+#40 := (< #22 #39)
+#49 := (ite #40 #34 #48)
+#50 := (< #31 #49)
+#45 := (ite #40 #26 #44)
+#46 := (< #45 #31)
+#51 := (and #46 #50)
+#272 := (iff #51 #271)
+#269 := (iff #50 #266)
+#260 := (< #31 #255)
+#267 := (iff #260 #266)
+#268 := [rewrite]: #267
+#261 := (iff #50 #260)
+#258 := (= #49 #255)
+#252 := (ite #227 #34 #48)
+#256 := (= #252 #255)
+#257 := [rewrite]: #256
+#253 := (= #49 #252)
+#231 := (iff #40 #227)
+#232 := [rewrite]: #231
+#254 := [monotonicity #232]: #253
+#259 := [trans #254 #257]: #258
+#262 := [monotonicity #259]: #261
+#270 := [trans #262 #268]: #269
+#250 := (iff #46 #247)
+#241 := (< #236 #31)
+#248 := (iff #241 #247)
+#249 := [rewrite]: #248
+#242 := (iff #46 #241)
+#239 := (= #45 #236)
+#233 := (ite #227 #26 #44)
+#237 := (= #233 #236)
+#238 := [rewrite]: #237
+#234 := (= #45 #233)
+#235 := [monotonicity #232]: #234
+#240 := [trans #235 #238]: #239
+#243 := [monotonicity #240]: #242
+#251 := [trans #243 #249]: #250
+#273 := [monotonicity #251 #270]: #272
+#178 := [asserted]: #51
+#274 := [mp #178 #273]: #271
+#275 := [and-elim #274]: #247
+#196 := (* -1::real #31)
+#212 := (+ #26 #196)
+#213 := (<= #212 0::real)
+#214 := (not #213)
+#197 := (+ #28 #196)
+#195 := (>= #197 0::real)
+#193 := (not #195)
+#219 := (and #193 #214)
+#23 := (< #22 #22)
+#35 := (ite #23 #34 #26)
+#36 := (< #31 #35)
+#29 := (ite #23 #26 #28)
+#32 := (< #29 #31)
+#37 := (and #32 #36)
+#220 := (iff #37 #219)
+#217 := (iff #36 #214)
+#209 := (< #31 #26)
+#215 := (iff #209 #214)
+#216 := [rewrite]: #215
+#210 := (iff #36 #209)
+#207 := (= #35 #26)
+#202 := (ite false #34 #26)
+#205 := (= #202 #26)
+#206 := [rewrite]: #205
+#203 := (= #35 #202)
+#180 := (iff #23 false)
+#181 := [rewrite]: #180
+#204 := [monotonicity #181]: #203
+#208 := [trans #204 #206]: #207
+#211 := [monotonicity #208]: #210
+#218 := [trans #211 #216]: #217
+#200 := (iff #32 #193)
+#189 := (< #28 #31)
+#198 := (iff #189 #193)
+#199 := [rewrite]: #198
+#190 := (iff #32 #189)
+#187 := (= #29 #28)
+#182 := (ite false #26 #28)
+#185 := (= #182 #28)
+#186 := [rewrite]: #185
+#183 := (= #29 #182)
+#184 := [monotonicity #181]: #183
+#188 := [trans #184 #186]: #187
+#191 := [monotonicity #188]: #190
+#201 := [trans #191 #199]: #200
+#221 := [monotonicity #201 #218]: #220
+#177 := [asserted]: #37
+#222 := [mp #177 #221]: #219
+#224 := [and-elim #222]: #214
+[th-lemma #224 #275 #689]: false
+unsat
+00b949278548f13329cf92f11a0ad2161fa40793 907 0
#2 := false
#299 := 0::real
decl uf_1 :: (-> T3 T2 real)
@@ -906,856 +1755,7 @@
#1366 := [unit-resolution #1233 #1365]: #1231
[th-lemma #1366 #1364 #1362]: false
unsat
-NQHwTeL311Tq3wf2s5BReA 419 0
-#2 := false
-#194 := 0::real
-decl uf_4 :: (-> T2 T3 real)
-decl uf_6 :: (-> T1 T3)
-decl uf_3 :: T1
-#21 := uf_3
-#25 := (uf_6 uf_3)
-decl uf_5 :: T2
-#24 := uf_5
-#26 := (uf_4 uf_5 #25)
-decl uf_7 :: T2
-#27 := uf_7
-#28 := (uf_4 uf_7 #25)
-decl uf_10 :: T1
-#38 := uf_10
-#42 := (uf_6 uf_10)
-decl uf_9 :: T2
-#33 := uf_9
-#43 := (uf_4 uf_9 #42)
-#41 := (= uf_3 uf_10)
-#44 := (ite #41 #43 #28)
-#9 := 0::int
-decl uf_2 :: (-> T1 int)
-#39 := (uf_2 uf_10)
-#226 := -1::int
-#229 := (* -1::int #39)
-#22 := (uf_2 uf_3)
-#230 := (+ #22 #229)
-#228 := (>= #230 0::int)
-#236 := (ite #228 #44 #26)
-#192 := -1::real
-#244 := (* -1::real #236)
-#642 := (+ #26 #244)
-#643 := (<= #642 0::real)
-#567 := (= #26 #236)
-#227 := (not #228)
-decl uf_1 :: (-> int T1)
-#593 := (uf_1 #39)
-#660 := (= #593 uf_10)
-#594 := (= uf_10 #593)
-#4 := (:var 0 T1)
-#5 := (uf_2 #4)
-#546 := (pattern #5)
-#6 := (uf_1 #5)
-#93 := (= #4 #6)
-#547 := (forall (vars (?x1 T1)) (:pat #546) #93)
-#96 := (forall (vars (?x1 T1)) #93)
-#550 := (iff #96 #547)
-#548 := (iff #93 #93)
-#549 := [refl]: #548
-#551 := [quant-intro #549]: #550
-#448 := (~ #96 #96)
-#450 := (~ #93 #93)
-#451 := [refl]: #450
-#449 := [nnf-pos #451]: #448
-#7 := (= #6 #4)
-#8 := (forall (vars (?x1 T1)) #7)
-#97 := (iff #8 #96)
-#94 := (iff #7 #93)
-#95 := [rewrite]: #94
-#98 := [quant-intro #95]: #97
-#92 := [asserted]: #8
-#101 := [mp #92 #98]: #96
-#446 := [mp~ #101 #449]: #96
-#552 := [mp #446 #551]: #547
-#595 := (not #547)
-#600 := (or #595 #594)
-#601 := [quant-inst]: #600
-#654 := [unit-resolution #601 #552]: #594
-#680 := [symm #654]: #660
-#681 := (= uf_3 #593)
-#591 := (uf_1 #22)
-#658 := (= #591 #593)
-#656 := (= #593 #591)
-#652 := (= #39 #22)
-#647 := (= #22 #39)
-#290 := (<= #230 0::int)
-#70 := (<= #22 #39)
-#388 := (iff #70 #290)
-#389 := [rewrite]: #388
-#341 := [asserted]: #70
-#390 := [mp #341 #389]: #290
-#646 := [hypothesis]: #228
-#648 := [th-lemma #646 #390]: #647
-#653 := [symm #648]: #652
-#657 := [monotonicity #653]: #656
-#659 := [symm #657]: #658
-#592 := (= uf_3 #591)
-#596 := (or #595 #592)
-#597 := [quant-inst]: #596
-#655 := [unit-resolution #597 #552]: #592
-#682 := [trans #655 #659]: #681
-#683 := [trans #682 #680]: #41
-#570 := (not #41)
-decl uf_11 :: T2
-#47 := uf_11
-#59 := (uf_4 uf_11 #42)
-#278 := (ite #41 #26 #59)
-#459 := (* -1::real #278)
-#637 := (+ #26 #459)
-#639 := (>= #637 0::real)
-#585 := (= #26 #278)
-#661 := [hypothesis]: #41
-#587 := (or #570 #585)
-#588 := [def-axiom]: #587
-#662 := [unit-resolution #588 #661]: #585
-#663 := (not #585)
-#664 := (or #663 #639)
-#665 := [th-lemma]: #664
-#666 := [unit-resolution #665 #662]: #639
-decl uf_8 :: T2
-#30 := uf_8
-#56 := (uf_4 uf_8 #42)
-#357 := (* -1::real #56)
-#358 := (+ #43 #357)
-#356 := (>= #358 0::real)
-#355 := (not #356)
-#374 := (* -1::real #59)
-#375 := (+ #56 #374)
-#373 := (>= #375 0::real)
-#376 := (not #373)
-#381 := (and #355 #376)
-#64 := (< #39 #39)
-#67 := (ite #64 #43 #59)
-#68 := (< #56 #67)
-#53 := (uf_4 uf_5 #42)
-#65 := (ite #64 #53 #43)
-#66 := (< #65 #56)
-#69 := (and #66 #68)
-#382 := (iff #69 #381)
-#379 := (iff #68 #376)
-#370 := (< #56 #59)
-#377 := (iff #370 #376)
-#378 := [rewrite]: #377
-#371 := (iff #68 #370)
-#368 := (= #67 #59)
-#363 := (ite false #43 #59)
-#366 := (= #363 #59)
-#367 := [rewrite]: #366
-#364 := (= #67 #363)
-#343 := (iff #64 false)
-#344 := [rewrite]: #343
-#365 := [monotonicity #344]: #364
-#369 := [trans #365 #367]: #368
-#372 := [monotonicity #369]: #371
-#380 := [trans #372 #378]: #379
-#361 := (iff #66 #355)
-#352 := (< #43 #56)
-#359 := (iff #352 #355)
-#360 := [rewrite]: #359
-#353 := (iff #66 #352)
-#350 := (= #65 #43)
-#345 := (ite false #53 #43)
-#348 := (= #345 #43)
-#349 := [rewrite]: #348
-#346 := (= #65 #345)
-#347 := [monotonicity #344]: #346
-#351 := [trans #347 #349]: #350
-#354 := [monotonicity #351]: #353
-#362 := [trans #354 #360]: #361
-#383 := [monotonicity #362 #380]: #382
-#340 := [asserted]: #69
-#384 := [mp #340 #383]: #381
-#385 := [and-elim #384]: #355
-#394 := (* -1::real #53)
-#395 := (+ #43 #394)
-#393 := (>= #395 0::real)
-#54 := (uf_4 uf_7 #42)
-#402 := (* -1::real #54)
-#403 := (+ #53 #402)
-#401 := (>= #403 0::real)
-#397 := (+ #43 #374)
-#398 := (<= #397 0::real)
-#412 := (and #393 #398 #401)
-#73 := (<= #43 #59)
-#72 := (<= #53 #43)
-#74 := (and #72 #73)
-#71 := (<= #54 #53)
-#75 := (and #71 #74)
-#415 := (iff #75 #412)
-#406 := (and #393 #398)
-#409 := (and #401 #406)
-#413 := (iff #409 #412)
-#414 := [rewrite]: #413
-#410 := (iff #75 #409)
-#407 := (iff #74 #406)
-#399 := (iff #73 #398)
-#400 := [rewrite]: #399
-#392 := (iff #72 #393)
-#396 := [rewrite]: #392
-#408 := [monotonicity #396 #400]: #407
-#404 := (iff #71 #401)
-#405 := [rewrite]: #404
-#411 := [monotonicity #405 #408]: #410
-#416 := [trans #411 #414]: #415
-#342 := [asserted]: #75
-#417 := [mp #342 #416]: #412
-#418 := [and-elim #417]: #393
-#650 := (+ #26 #394)
-#651 := (<= #650 0::real)
-#649 := (= #26 #53)
-#671 := (= #53 #26)
-#669 := (= #42 #25)
-#667 := (= #25 #42)
-#668 := [monotonicity #661]: #667
-#670 := [symm #668]: #669
-#672 := [monotonicity #670]: #671
-#673 := [symm #672]: #649
-#674 := (not #649)
-#675 := (or #674 #651)
-#676 := [th-lemma]: #675
-#677 := [unit-resolution #676 #673]: #651
-#462 := (+ #56 #459)
-#465 := (>= #462 0::real)
-#438 := (not #465)
-#316 := (ite #290 #278 #43)
-#326 := (* -1::real #316)
-#327 := (+ #56 #326)
-#325 := (>= #327 0::real)
-#324 := (not #325)
-#439 := (iff #324 #438)
-#466 := (iff #325 #465)
-#463 := (= #327 #462)
-#460 := (= #326 #459)
-#457 := (= #316 #278)
-#1 := true
-#452 := (ite true #278 #43)
-#455 := (= #452 #278)
-#456 := [rewrite]: #455
-#453 := (= #316 #452)
-#444 := (iff #290 true)
-#445 := [iff-true #390]: #444
-#454 := [monotonicity #445]: #453
-#458 := [trans #454 #456]: #457
-#461 := [monotonicity #458]: #460
-#464 := [monotonicity #461]: #463
-#467 := [monotonicity #464]: #466
-#468 := [monotonicity #467]: #439
-#297 := (ite #290 #54 #53)
-#305 := (* -1::real #297)
-#306 := (+ #56 #305)
-#307 := (<= #306 0::real)
-#308 := (not #307)
-#332 := (and #308 #324)
-#58 := (= uf_10 uf_3)
-#60 := (ite #58 #26 #59)
-#52 := (< #39 #22)
-#61 := (ite #52 #43 #60)
-#62 := (< #56 #61)
-#55 := (ite #52 #53 #54)
-#57 := (< #55 #56)
-#63 := (and #57 #62)
-#335 := (iff #63 #332)
-#281 := (ite #52 #43 #278)
-#284 := (< #56 #281)
-#287 := (and #57 #284)
-#333 := (iff #287 #332)
-#330 := (iff #284 #324)
-#321 := (< #56 #316)
-#328 := (iff #321 #324)
-#329 := [rewrite]: #328
-#322 := (iff #284 #321)
-#319 := (= #281 #316)
-#291 := (not #290)
-#313 := (ite #291 #43 #278)
-#317 := (= #313 #316)
-#318 := [rewrite]: #317
-#314 := (= #281 #313)
-#292 := (iff #52 #291)
-#293 := [rewrite]: #292
-#315 := [monotonicity #293]: #314
-#320 := [trans #315 #318]: #319
-#323 := [monotonicity #320]: #322
-#331 := [trans #323 #329]: #330
-#311 := (iff #57 #308)
-#302 := (< #297 #56)
-#309 := (iff #302 #308)
-#310 := [rewrite]: #309
-#303 := (iff #57 #302)
-#300 := (= #55 #297)
-#294 := (ite #291 #53 #54)
-#298 := (= #294 #297)
-#299 := [rewrite]: #298
-#295 := (= #55 #294)
-#296 := [monotonicity #293]: #295
-#301 := [trans #296 #299]: #300
-#304 := [monotonicity #301]: #303
-#312 := [trans #304 #310]: #311
-#334 := [monotonicity #312 #331]: #333
-#288 := (iff #63 #287)
-#285 := (iff #62 #284)
-#282 := (= #61 #281)
-#279 := (= #60 #278)
-#225 := (iff #58 #41)
-#277 := [rewrite]: #225
-#280 := [monotonicity #277]: #279
-#283 := [monotonicity #280]: #282
-#286 := [monotonicity #283]: #285
-#289 := [monotonicity #286]: #288
-#336 := [trans #289 #334]: #335
-#179 := [asserted]: #63
-#337 := [mp #179 #336]: #332
-#339 := [and-elim #337]: #324
-#469 := [mp #339 #468]: #438
-#678 := [th-lemma #469 #677 #418 #385 #666]: false
-#679 := [lemma #678]: #570
-#684 := [unit-resolution #679 #683]: false
-#685 := [lemma #684]: #227
-#577 := (or #228 #567)
-#578 := [def-axiom]: #577
-#645 := [unit-resolution #578 #685]: #567
-#686 := (not #567)
-#687 := (or #686 #643)
-#688 := [th-lemma]: #687
-#689 := [unit-resolution #688 #645]: #643
-#31 := (uf_4 uf_8 #25)
-#245 := (+ #31 #244)
-#246 := (<= #245 0::real)
-#247 := (not #246)
-#34 := (uf_4 uf_9 #25)
-#48 := (uf_4 uf_11 #25)
-#255 := (ite #228 #48 #34)
-#264 := (* -1::real #255)
-#265 := (+ #31 #264)
-#263 := (>= #265 0::real)
-#266 := (not #263)
-#271 := (and #247 #266)
-#40 := (< #22 #39)
-#49 := (ite #40 #34 #48)
-#50 := (< #31 #49)
-#45 := (ite #40 #26 #44)
-#46 := (< #45 #31)
-#51 := (and #46 #50)
-#272 := (iff #51 #271)
-#269 := (iff #50 #266)
-#260 := (< #31 #255)
-#267 := (iff #260 #266)
-#268 := [rewrite]: #267
-#261 := (iff #50 #260)
-#258 := (= #49 #255)
-#252 := (ite #227 #34 #48)
-#256 := (= #252 #255)
-#257 := [rewrite]: #256
-#253 := (= #49 #252)
-#231 := (iff #40 #227)
-#232 := [rewrite]: #231
-#254 := [monotonicity #232]: #253
-#259 := [trans #254 #257]: #258
-#262 := [monotonicity #259]: #261
-#270 := [trans #262 #268]: #269
-#250 := (iff #46 #247)
-#241 := (< #236 #31)
-#248 := (iff #241 #247)
-#249 := [rewrite]: #248
-#242 := (iff #46 #241)
-#239 := (= #45 #236)
-#233 := (ite #227 #26 #44)
-#237 := (= #233 #236)
-#238 := [rewrite]: #237
-#234 := (= #45 #233)
-#235 := [monotonicity #232]: #234
-#240 := [trans #235 #238]: #239
-#243 := [monotonicity #240]: #242
-#251 := [trans #243 #249]: #250
-#273 := [monotonicity #251 #270]: #272
-#178 := [asserted]: #51
-#274 := [mp #178 #273]: #271
-#275 := [and-elim #274]: #247
-#196 := (* -1::real #31)
-#212 := (+ #26 #196)
-#213 := (<= #212 0::real)
-#214 := (not #213)
-#197 := (+ #28 #196)
-#195 := (>= #197 0::real)
-#193 := (not #195)
-#219 := (and #193 #214)
-#23 := (< #22 #22)
-#35 := (ite #23 #34 #26)
-#36 := (< #31 #35)
-#29 := (ite #23 #26 #28)
-#32 := (< #29 #31)
-#37 := (and #32 #36)
-#220 := (iff #37 #219)
-#217 := (iff #36 #214)
-#209 := (< #31 #26)
-#215 := (iff #209 #214)
-#216 := [rewrite]: #215
-#210 := (iff #36 #209)
-#207 := (= #35 #26)
-#202 := (ite false #34 #26)
-#205 := (= #202 #26)
-#206 := [rewrite]: #205
-#203 := (= #35 #202)
-#180 := (iff #23 false)
-#181 := [rewrite]: #180
-#204 := [monotonicity #181]: #203
-#208 := [trans #204 #206]: #207
-#211 := [monotonicity #208]: #210
-#218 := [trans #211 #216]: #217
-#200 := (iff #32 #193)
-#189 := (< #28 #31)
-#198 := (iff #189 #193)
-#199 := [rewrite]: #198
-#190 := (iff #32 #189)
-#187 := (= #29 #28)
-#182 := (ite false #26 #28)
-#185 := (= #182 #28)
-#186 := [rewrite]: #185
-#183 := (= #29 #182)
-#184 := [monotonicity #181]: #183
-#188 := [trans #184 #186]: #187
-#191 := [monotonicity #188]: #190
-#201 := [trans #191 #199]: #200
-#221 := [monotonicity #201 #218]: #220
-#177 := [asserted]: #37
-#222 := [mp #177 #221]: #219
-#224 := [and-elim #222]: #214
-[th-lemma #224 #275 #689]: false
-unsat
-NX/HT1QOfbspC2LtZNKpBA 428 0
-#2 := false
-decl uf_10 :: T1
-#38 := uf_10
-decl uf_3 :: T1
-#21 := uf_3
-#45 := (= uf_3 uf_10)
-decl uf_1 :: (-> int T1)
-decl uf_2 :: (-> T1 int)
-#39 := (uf_2 uf_10)
-#588 := (uf_1 #39)
-#686 := (= #588 uf_10)
-#589 := (= uf_10 #588)
-#4 := (:var 0 T1)
-#5 := (uf_2 #4)
-#541 := (pattern #5)
-#6 := (uf_1 #5)
-#93 := (= #4 #6)
-#542 := (forall (vars (?x1 T1)) (:pat #541) #93)
-#96 := (forall (vars (?x1 T1)) #93)
-#545 := (iff #96 #542)
-#543 := (iff #93 #93)
-#544 := [refl]: #543
-#546 := [quant-intro #544]: #545
-#454 := (~ #96 #96)
-#456 := (~ #93 #93)
-#457 := [refl]: #456
-#455 := [nnf-pos #457]: #454
-#7 := (= #6 #4)
-#8 := (forall (vars (?x1 T1)) #7)
-#97 := (iff #8 #96)
-#94 := (iff #7 #93)
-#95 := [rewrite]: #94
-#98 := [quant-intro #95]: #97
-#92 := [asserted]: #8
-#101 := [mp #92 #98]: #96
-#452 := [mp~ #101 #455]: #96
-#547 := [mp #452 #546]: #542
-#590 := (not #542)
-#595 := (or #590 #589)
-#596 := [quant-inst]: #595
-#680 := [unit-resolution #596 #547]: #589
-#687 := [symm #680]: #686
-#688 := (= uf_3 #588)
-#22 := (uf_2 uf_3)
-#586 := (uf_1 #22)
-#684 := (= #586 #588)
-#682 := (= #588 #586)
-#678 := (= #39 #22)
-#676 := (= #22 #39)
-#9 := 0::int
-#227 := -1::int
-#230 := (* -1::int #39)
-#231 := (+ #22 #230)
-#296 := (<= #231 0::int)
-#70 := (<= #22 #39)
-#393 := (iff #70 #296)
-#394 := [rewrite]: #393
-#347 := [asserted]: #70
-#395 := [mp #347 #394]: #296
-#229 := (>= #231 0::int)
-decl uf_4 :: (-> T2 T3 real)
-decl uf_6 :: (-> T1 T3)
-#25 := (uf_6 uf_3)
-decl uf_7 :: T2
-#27 := uf_7
-#28 := (uf_4 uf_7 #25)
-decl uf_9 :: T2
-#33 := uf_9
-#34 := (uf_4 uf_9 #25)
-#46 := (uf_6 uf_10)
-decl uf_5 :: T2
-#24 := uf_5
-#47 := (uf_4 uf_5 #46)
-#48 := (ite #45 #47 #34)
-#256 := (ite #229 #48 #28)
-#568 := (= #28 #256)
-#648 := (not #568)
-#194 := 0::real
-#192 := -1::real
-#265 := (* -1::real #256)
-#640 := (+ #28 #265)
-#642 := (>= #640 0::real)
-#645 := (not #642)
-#643 := [hypothesis]: #642
-decl uf_8 :: T2
-#30 := uf_8
-#31 := (uf_4 uf_8 #25)
-#266 := (+ #31 #265)
-#264 := (>= #266 0::real)
-#267 := (not #264)
-#26 := (uf_4 uf_5 #25)
-decl uf_11 :: T2
-#41 := uf_11
-#42 := (uf_4 uf_11 #25)
-#237 := (ite #229 #42 #26)
-#245 := (* -1::real #237)
-#246 := (+ #31 #245)
-#247 := (<= #246 0::real)
-#248 := (not #247)
-#272 := (and #248 #267)
-#40 := (< #22 #39)
-#49 := (ite #40 #28 #48)
-#50 := (< #31 #49)
-#43 := (ite #40 #26 #42)
-#44 := (< #43 #31)
-#51 := (and #44 #50)
-#273 := (iff #51 #272)
-#270 := (iff #50 #267)
-#261 := (< #31 #256)
-#268 := (iff #261 #267)
-#269 := [rewrite]: #268
-#262 := (iff #50 #261)
-#259 := (= #49 #256)
-#228 := (not #229)
-#253 := (ite #228 #28 #48)
-#257 := (= #253 #256)
-#258 := [rewrite]: #257
-#254 := (= #49 #253)
-#232 := (iff #40 #228)
-#233 := [rewrite]: #232
-#255 := [monotonicity #233]: #254
-#260 := [trans #255 #258]: #259
-#263 := [monotonicity #260]: #262
-#271 := [trans #263 #269]: #270
-#251 := (iff #44 #248)
-#242 := (< #237 #31)
-#249 := (iff #242 #248)
-#250 := [rewrite]: #249
-#243 := (iff #44 #242)
-#240 := (= #43 #237)
-#234 := (ite #228 #26 #42)
-#238 := (= #234 #237)
-#239 := [rewrite]: #238
-#235 := (= #43 #234)
-#236 := [monotonicity #233]: #235
-#241 := [trans #236 #239]: #240
-#244 := [monotonicity #241]: #243
-#252 := [trans #244 #250]: #251
-#274 := [monotonicity #252 #271]: #273
-#178 := [asserted]: #51
-#275 := [mp #178 #274]: #272
-#277 := [and-elim #275]: #267
-#196 := (* -1::real #31)
-#197 := (+ #28 #196)
-#195 := (>= #197 0::real)
-#193 := (not #195)
-#213 := (* -1::real #34)
-#214 := (+ #31 #213)
-#212 := (>= #214 0::real)
-#215 := (not #212)
-#220 := (and #193 #215)
-#23 := (< #22 #22)
-#35 := (ite #23 #28 #34)
-#36 := (< #31 #35)
-#29 := (ite #23 #26 #28)
-#32 := (< #29 #31)
-#37 := (and #32 #36)
-#221 := (iff #37 #220)
-#218 := (iff #36 #215)
-#209 := (< #31 #34)
-#216 := (iff #209 #215)
-#217 := [rewrite]: #216
-#210 := (iff #36 #209)
-#207 := (= #35 #34)
-#202 := (ite false #28 #34)
-#205 := (= #202 #34)
-#206 := [rewrite]: #205
-#203 := (= #35 #202)
-#180 := (iff #23 false)
-#181 := [rewrite]: #180
-#204 := [monotonicity #181]: #203
-#208 := [trans #204 #206]: #207
-#211 := [monotonicity #208]: #210
-#219 := [trans #211 #217]: #218
-#200 := (iff #32 #193)
-#189 := (< #28 #31)
-#198 := (iff #189 #193)
-#199 := [rewrite]: #198
-#190 := (iff #32 #189)
-#187 := (= #29 #28)
-#182 := (ite false #26 #28)
-#185 := (= #182 #28)
-#186 := [rewrite]: #185
-#183 := (= #29 #182)
-#184 := [monotonicity #181]: #183
-#188 := [trans #184 #186]: #187
-#191 := [monotonicity #188]: #190
-#201 := [trans #191 #199]: #200
-#222 := [monotonicity #201 #219]: #221
-#177 := [asserted]: #37
-#223 := [mp #177 #222]: #220
-#224 := [and-elim #223]: #193
-#644 := [th-lemma #224 #277 #643]: false
-#646 := [lemma #644]: #645
-#647 := [hypothesis]: #568
-#649 := (or #648 #642)
-#650 := [th-lemma]: #649
-#651 := [unit-resolution #650 #647 #646]: false
-#652 := [lemma #651]: #648
-#578 := (or #229 #568)
-#579 := [def-axiom]: #578
-#675 := [unit-resolution #579 #652]: #229
-#677 := [th-lemma #675 #395]: #676
-#679 := [symm #677]: #678
-#683 := [monotonicity #679]: #682
-#685 := [symm #683]: #684
-#587 := (= uf_3 #586)
-#591 := (or #590 #587)
-#592 := [quant-inst]: #591
-#681 := [unit-resolution #592 #547]: #587
-#689 := [trans #681 #685]: #688
-#690 := [trans #689 #687]: #45
-#571 := (not #45)
-#54 := (uf_4 uf_11 #46)
-#279 := (ite #45 #28 #54)
-#465 := (* -1::real #279)
-#632 := (+ #28 #465)
-#633 := (<= #632 0::real)
-#580 := (= #28 #279)
-#656 := [hypothesis]: #45
-#582 := (or #571 #580)
-#583 := [def-axiom]: #582
-#657 := [unit-resolution #583 #656]: #580
-#658 := (not #580)
-#659 := (or #658 #633)
-#660 := [th-lemma]: #659
-#661 := [unit-resolution #660 #657]: #633
-#57 := (uf_4 uf_8 #46)
-#363 := (* -1::real #57)
-#379 := (+ #47 #363)
-#380 := (<= #379 0::real)
-#381 := (not #380)
-#364 := (+ #54 #363)
-#362 := (>= #364 0::real)
-#361 := (not #362)
-#386 := (and #361 #381)
-#59 := (uf_4 uf_7 #46)
-#64 := (< #39 #39)
-#67 := (ite #64 #59 #47)
-#68 := (< #57 #67)
-#65 := (ite #64 #47 #54)
-#66 := (< #65 #57)
-#69 := (and #66 #68)
-#387 := (iff #69 #386)
-#384 := (iff #68 #381)
-#376 := (< #57 #47)
-#382 := (iff #376 #381)
-#383 := [rewrite]: #382
-#377 := (iff #68 #376)
-#374 := (= #67 #47)
-#369 := (ite false #59 #47)
-#372 := (= #369 #47)
-#373 := [rewrite]: #372
-#370 := (= #67 #369)
-#349 := (iff #64 false)
-#350 := [rewrite]: #349
-#371 := [monotonicity #350]: #370
-#375 := [trans #371 #373]: #374
-#378 := [monotonicity #375]: #377
-#385 := [trans #378 #383]: #384
-#367 := (iff #66 #361)
-#358 := (< #54 #57)
-#365 := (iff #358 #361)
-#366 := [rewrite]: #365
-#359 := (iff #66 #358)
-#356 := (= #65 #54)
-#351 := (ite false #47 #54)
-#354 := (= #351 #54)
-#355 := [rewrite]: #354
-#352 := (= #65 #351)
-#353 := [monotonicity #350]: #352
-#357 := [trans #353 #355]: #356
-#360 := [monotonicity #357]: #359
-#368 := [trans #360 #366]: #367
-#388 := [monotonicity #368 #385]: #387
-#346 := [asserted]: #69
-#389 := [mp #346 #388]: #386
-#391 := [and-elim #389]: #381
-#397 := (* -1::real #59)
-#398 := (+ #47 #397)
-#399 := (<= #398 0::real)
-#409 := (* -1::real #54)
-#410 := (+ #47 #409)
-#408 := (>= #410 0::real)
-#60 := (uf_4 uf_9 #46)
-#402 := (* -1::real #60)
-#403 := (+ #59 #402)
-#404 := (<= #403 0::real)
-#418 := (and #399 #404 #408)
-#73 := (<= #59 #60)
-#72 := (<= #47 #59)
-#74 := (and #72 #73)
-#71 := (<= #54 #47)
-#75 := (and #71 #74)
-#421 := (iff #75 #418)
-#412 := (and #399 #404)
-#415 := (and #408 #412)
-#419 := (iff #415 #418)
-#420 := [rewrite]: #419
-#416 := (iff #75 #415)
-#413 := (iff #74 #412)
-#405 := (iff #73 #404)
-#406 := [rewrite]: #405
-#400 := (iff #72 #399)
-#401 := [rewrite]: #400
-#414 := [monotonicity #401 #406]: #413
-#407 := (iff #71 #408)
-#411 := [rewrite]: #407
-#417 := [monotonicity #411 #414]: #416
-#422 := [trans #417 #420]: #421
-#348 := [asserted]: #75
-#423 := [mp #348 #422]: #418
-#424 := [and-elim #423]: #399
-#637 := (+ #28 #397)
-#639 := (>= #637 0::real)
-#636 := (= #28 #59)
-#666 := (= #59 #28)
-#664 := (= #46 #25)
-#662 := (= #25 #46)
-#663 := [monotonicity #656]: #662
-#665 := [symm #663]: #664
-#667 := [monotonicity #665]: #666
-#668 := [symm #667]: #636
-#669 := (not #636)
-#670 := (or #669 #639)
-#671 := [th-lemma]: #670
-#672 := [unit-resolution #671 #668]: #639
-#468 := (+ #57 #465)
-#471 := (<= #468 0::real)
-#444 := (not #471)
-#322 := (ite #296 #279 #47)
-#330 := (* -1::real #322)
-#331 := (+ #57 #330)
-#332 := (<= #331 0::real)
-#333 := (not #332)
-#445 := (iff #333 #444)
-#472 := (iff #332 #471)
-#469 := (= #331 #468)
-#466 := (= #330 #465)
-#463 := (= #322 #279)
-#1 := true
-#458 := (ite true #279 #47)
-#461 := (= #458 #279)
-#462 := [rewrite]: #461
-#459 := (= #322 #458)
-#450 := (iff #296 true)
-#451 := [iff-true #395]: #450
-#460 := [monotonicity #451]: #459
-#464 := [trans #460 #462]: #463
-#467 := [monotonicity #464]: #466
-#470 := [monotonicity #467]: #469
-#473 := [monotonicity #470]: #472
-#474 := [monotonicity #473]: #445
-#303 := (ite #296 #60 #59)
-#313 := (* -1::real #303)
-#314 := (+ #57 #313)
-#312 := (>= #314 0::real)
-#311 := (not #312)
-#338 := (and #311 #333)
-#52 := (< #39 #22)
-#61 := (ite #52 #59 #60)
-#62 := (< #57 #61)
-#53 := (= uf_10 uf_3)
-#55 := (ite #53 #28 #54)
-#56 := (ite #52 #47 #55)
-#58 := (< #56 #57)
-#63 := (and #58 #62)
-#341 := (iff #63 #338)
-#282 := (ite #52 #47 #279)
-#285 := (< #282 #57)
-#291 := (and #62 #285)
-#339 := (iff #291 #338)
-#336 := (iff #285 #333)
-#327 := (< #322 #57)
-#334 := (iff #327 #333)
-#335 := [rewrite]: #334
-#328 := (iff #285 #327)
-#325 := (= #282 #322)
-#297 := (not #296)
-#319 := (ite #297 #47 #279)
-#323 := (= #319 #322)
-#324 := [rewrite]: #323
-#320 := (= #282 #319)
-#298 := (iff #52 #297)
-#299 := [rewrite]: #298
-#321 := [monotonicity #299]: #320
-#326 := [trans #321 #324]: #325
-#329 := [monotonicity #326]: #328
-#337 := [trans #329 #335]: #336
-#317 := (iff #62 #311)
-#308 := (< #57 #303)
-#315 := (iff #308 #311)
-#316 := [rewrite]: #315
-#309 := (iff #62 #308)
-#306 := (= #61 #303)
-#300 := (ite #297 #59 #60)
-#304 := (= #300 #303)
-#305 := [rewrite]: #304
-#301 := (= #61 #300)
-#302 := [monotonicity #299]: #301
-#307 := [trans #302 #305]: #306
-#310 := [monotonicity #307]: #309
-#318 := [trans #310 #316]: #317
-#340 := [monotonicity #318 #337]: #339
-#294 := (iff #63 #291)
-#288 := (and #285 #62)
-#292 := (iff #288 #291)
-#293 := [rewrite]: #292
-#289 := (iff #63 #288)
-#286 := (iff #58 #285)
-#283 := (= #56 #282)
-#280 := (= #55 #279)
-#226 := (iff #53 #45)
-#278 := [rewrite]: #226
-#281 := [monotonicity #278]: #280
-#284 := [monotonicity #281]: #283
-#287 := [monotonicity #284]: #286
-#290 := [monotonicity #287]: #289
-#295 := [trans #290 #293]: #294
-#342 := [trans #295 #340]: #341
-#179 := [asserted]: #63
-#343 := [mp #179 #342]: #338
-#345 := [and-elim #343]: #333
-#475 := [mp #345 #474]: #444
-#673 := [th-lemma #475 #672 #424 #391 #661]: false
-#674 := [lemma #673]: #571
-[unit-resolution #674 #690]: false
-unsat
-IL2powemHjRpCJYwmXFxyw 211 0
+f5cf3017773c1cf700f533cb220ba6c9eb22de56 211 0
#2 := false
#33 := 0::real
decl uf_11 :: (-> T5 T6 real)
@@ -1967,7 +1967,7 @@
#328 := [unit-resolution #319 #327]: #300
[th-lemma #326 #334 #195 #328 #187 #138]: false
unsat
-GX51o3DUO/UBS3eNP2P9kA 285 0
+ad41c2d9b185dd9b4fd0abefcdc55357d2105ed8 285 0
#2 := false
#7 := 0::real
decl uf_4 :: real
@@ -2253,7 +2253,7 @@
#286 := [lemma #284]: #285
[unit-resolution #286 #308 #305]: false
unsat
-cebG074uorSr8ODzgTmcKg 97 0
+076d8565ad4c4370b16a3b0c139e8c8e8fb58224 97 0
#2 := false
#18 := 0::real
decl uf_1 :: (-> T2 T1 real)
@@ -2351,7 +2351,7 @@
#173 := [mp #167 #172]: #165
[unit-resolution #173 #147 #76]: false
unsat
-DKRtrJ2XceCkITuNwNViRw 57 0
+8e6ea05b53a9ffb5a6dec689eb004bdb0c2a755b 57 0
#2 := false
#4 := 0::real
decl uf_1 :: (-> T2 real)
@@ -2409,7 +2409,7 @@
#269 := [quant-inst]: #268
[unit-resolution #269 #247 #274]: false
unsat
-97KJAJfUio+nGchEHWvgAw 91 0
+d278c38176edc1fb6492d54817f0a9092db653e6 91 0
#2 := false
#38 := 0::real
decl uf_1 :: (-> T1 T2 real)
@@ -2501,157 +2501,7 @@
#153 := [mp #147 #152]: #145
[unit-resolution #153 #129 #160]: false
unsat
-flJYbeWfe+t2l/zsRqdujA 149 0
-#2 := false
-#19 := 0::real
-decl uf_1 :: (-> T1 T2 real)
-decl uf_3 :: T2
-#5 := uf_3
-decl uf_4 :: T1
-#7 := uf_4
-#8 := (uf_1 uf_4 uf_3)
-#44 := -1::real
-#156 := (* -1::real #8)
-decl uf_2 :: T1
-#4 := uf_2
-#6 := (uf_1 uf_2 uf_3)
-#203 := (+ #6 #156)
-#205 := (>= #203 0::real)
-#9 := (= #6 #8)
-#40 := [asserted]: #9
-#208 := (not #9)
-#209 := (or #208 #205)
-#210 := [th-lemma]: #209
-#211 := [unit-resolution #210 #40]: #205
-decl uf_5 :: T1
-#12 := uf_5
-#22 := (uf_1 uf_5 uf_3)
-#160 := (* -1::real #22)
-#161 := (+ #6 #160)
-#207 := (>= #161 0::real)
-#222 := (not #207)
-#206 := (= #6 #22)
-#216 := (not #206)
-#62 := (= #8 #22)
-#70 := (not #62)
-#217 := (iff #70 #216)
-#214 := (iff #62 #206)
-#212 := (iff #206 #62)
-#213 := [monotonicity #40]: #212
-#215 := [symm #213]: #214
-#218 := [monotonicity #215]: #217
-#23 := (= #22 #8)
-#24 := (not #23)
-#71 := (iff #24 #70)
-#68 := (iff #23 #62)
-#69 := [rewrite]: #68
-#72 := [monotonicity #69]: #71
-#43 := [asserted]: #24
-#75 := [mp #43 #72]: #70
-#219 := [mp #75 #218]: #216
-#225 := (or #206 #222)
-#162 := (<= #161 0::real)
-#172 := (+ #8 #160)
-#173 := (>= #172 0::real)
-#178 := (not #173)
-#163 := (not #162)
-#181 := (or #163 #178)
-#184 := (not #181)
-#10 := (:var 0 T2)
-#15 := (uf_1 uf_4 #10)
-#149 := (pattern #15)
-#13 := (uf_1 uf_5 #10)
-#148 := (pattern #13)
-#11 := (uf_1 uf_2 #10)
-#147 := (pattern #11)
-#50 := (* -1::real #15)
-#51 := (+ #13 #50)
-#52 := (<= #51 0::real)
-#76 := (not #52)
-#45 := (* -1::real #13)
-#46 := (+ #11 #45)
-#47 := (<= #46 0::real)
-#78 := (not #47)
-#73 := (or #78 #76)
-#83 := (not #73)
-#150 := (forall (vars (?x1 T2)) (:pat #147 #148 #149) #83)
-#86 := (forall (vars (?x1 T2)) #83)
-#153 := (iff #86 #150)
-#151 := (iff #83 #83)
-#152 := [refl]: #151
-#154 := [quant-intro #152]: #153
-#55 := (and #47 #52)
-#58 := (forall (vars (?x1 T2)) #55)
-#87 := (iff #58 #86)
-#84 := (iff #55 #83)
-#85 := [rewrite]: #84
-#88 := [quant-intro #85]: #87
-#79 := (~ #58 #58)
-#81 := (~ #55 #55)
-#82 := [refl]: #81
-#80 := [nnf-pos #82]: #79
-#16 := (<= #13 #15)
-#14 := (<= #11 #13)
-#17 := (and #14 #16)
-#18 := (forall (vars (?x1 T2)) #17)
-#59 := (iff #18 #58)
-#56 := (iff #17 #55)
-#53 := (iff #16 #52)
-#54 := [rewrite]: #53
-#48 := (iff #14 #47)
-#49 := [rewrite]: #48
-#57 := [monotonicity #49 #54]: #56
-#60 := [quant-intro #57]: #59
-#41 := [asserted]: #18
-#61 := [mp #41 #60]: #58
-#77 := [mp~ #61 #80]: #58
-#89 := [mp #77 #88]: #86
-#155 := [mp #89 #154]: #150
-#187 := (not #150)
-#188 := (or #187 #184)
-#157 := (+ #22 #156)
-#158 := (<= #157 0::real)
-#159 := (not #158)
-#164 := (or #163 #159)
-#165 := (not #164)
-#189 := (or #187 #165)
-#191 := (iff #189 #188)
-#193 := (iff #188 #188)
-#194 := [rewrite]: #193
-#185 := (iff #165 #184)
-#182 := (iff #164 #181)
-#179 := (iff #159 #178)
-#176 := (iff #158 #173)
-#166 := (+ #156 #22)
-#169 := (<= #166 0::real)
-#174 := (iff #169 #173)
-#175 := [rewrite]: #174
-#170 := (iff #158 #169)
-#167 := (= #157 #166)
-#168 := [rewrite]: #167
-#171 := [monotonicity #168]: #170
-#177 := [trans #171 #175]: #176
-#180 := [monotonicity #177]: #179
-#183 := [monotonicity #180]: #182
-#186 := [monotonicity #183]: #185
-#192 := [monotonicity #186]: #191
-#195 := [trans #192 #194]: #191
-#190 := [quant-inst]: #189
-#196 := [mp #190 #195]: #188
-#220 := [unit-resolution #196 #155]: #184
-#197 := (or #181 #162)
-#198 := [def-axiom]: #197
-#221 := [unit-resolution #198 #220]: #162
-#223 := (or #206 #163 #222)
-#224 := [th-lemma]: #223
-#226 := [unit-resolution #224 #221]: #225
-#227 := [unit-resolution #226 #219]: #222
-#199 := (or #181 #173)
-#200 := [def-axiom]: #199
-#228 := [unit-resolution #200 #220]: #173
-[th-lemma #228 #227 #211]: false
-unsat
-rbrrQuQfaijtLkQizgEXnQ 222 0
+38b5e95092f91070fab780891ebf7b83d5f95757 222 0
#2 := false
#4 := 0::real
decl uf_2 :: (-> T2 T1 real)
@@ -2874,7 +2724,7 @@
#308 := [th-lemma]: #307
[unit-resolution #308 #305 #290]: false
unsat
-hwh3oeLAWt56hnKIa8Wuow 248 0
+4968b0a0840255d92fbecd59ed4dac302603b2bd 248 0
#2 := false
#4 := 0::real
decl uf_2 :: (-> T2 T1 real)
@@ -3123,7 +2973,157 @@
#380 := [th-lemma]: #379
[unit-resolution #380 #377 #363]: false
unsat
-WdMJH3tkMv/rps8y9Ukq5Q 86 0
+a3b02dd75d8b2261f1b7ef7215268d84fd25e972 149 0
+#2 := false
+#19 := 0::real
+decl uf_1 :: (-> T1 T2 real)
+decl uf_3 :: T2
+#5 := uf_3
+decl uf_4 :: T1
+#7 := uf_4
+#8 := (uf_1 uf_4 uf_3)
+#44 := -1::real
+#156 := (* -1::real #8)
+decl uf_2 :: T1
+#4 := uf_2
+#6 := (uf_1 uf_2 uf_3)
+#203 := (+ #6 #156)
+#205 := (>= #203 0::real)
+#9 := (= #6 #8)
+#40 := [asserted]: #9
+#208 := (not #9)
+#209 := (or #208 #205)
+#210 := [th-lemma]: #209
+#211 := [unit-resolution #210 #40]: #205
+decl uf_5 :: T1
+#12 := uf_5
+#22 := (uf_1 uf_5 uf_3)
+#160 := (* -1::real #22)
+#161 := (+ #6 #160)
+#207 := (>= #161 0::real)
+#222 := (not #207)
+#206 := (= #6 #22)
+#216 := (not #206)
+#62 := (= #8 #22)
+#70 := (not #62)
+#217 := (iff #70 #216)
+#214 := (iff #62 #206)
+#212 := (iff #206 #62)
+#213 := [monotonicity #40]: #212
+#215 := [symm #213]: #214
+#218 := [monotonicity #215]: #217
+#23 := (= #22 #8)
+#24 := (not #23)
+#71 := (iff #24 #70)
+#68 := (iff #23 #62)
+#69 := [rewrite]: #68
+#72 := [monotonicity #69]: #71
+#43 := [asserted]: #24
+#75 := [mp #43 #72]: #70
+#219 := [mp #75 #218]: #216
+#225 := (or #206 #222)
+#162 := (<= #161 0::real)
+#172 := (+ #8 #160)
+#173 := (>= #172 0::real)
+#178 := (not #173)
+#163 := (not #162)
+#181 := (or #163 #178)
+#184 := (not #181)
+#10 := (:var 0 T2)
+#15 := (uf_1 uf_4 #10)
+#149 := (pattern #15)
+#13 := (uf_1 uf_5 #10)
+#148 := (pattern #13)
+#11 := (uf_1 uf_2 #10)
+#147 := (pattern #11)
+#50 := (* -1::real #15)
+#51 := (+ #13 #50)
+#52 := (<= #51 0::real)
+#76 := (not #52)
+#45 := (* -1::real #13)
+#46 := (+ #11 #45)
+#47 := (<= #46 0::real)
+#78 := (not #47)
+#73 := (or #78 #76)
+#83 := (not #73)
+#150 := (forall (vars (?x1 T2)) (:pat #147 #148 #149) #83)
+#86 := (forall (vars (?x1 T2)) #83)
+#153 := (iff #86 #150)
+#151 := (iff #83 #83)
+#152 := [refl]: #151
+#154 := [quant-intro #152]: #153
+#55 := (and #47 #52)
+#58 := (forall (vars (?x1 T2)) #55)
+#87 := (iff #58 #86)
+#84 := (iff #55 #83)
+#85 := [rewrite]: #84
+#88 := [quant-intro #85]: #87
+#79 := (~ #58 #58)
+#81 := (~ #55 #55)
+#82 := [refl]: #81
+#80 := [nnf-pos #82]: #79
+#16 := (<= #13 #15)
+#14 := (<= #11 #13)
+#17 := (and #14 #16)
+#18 := (forall (vars (?x1 T2)) #17)
+#59 := (iff #18 #58)
+#56 := (iff #17 #55)
+#53 := (iff #16 #52)
+#54 := [rewrite]: #53
+#48 := (iff #14 #47)
+#49 := [rewrite]: #48
+#57 := [monotonicity #49 #54]: #56
+#60 := [quant-intro #57]: #59
+#41 := [asserted]: #18
+#61 := [mp #41 #60]: #58
+#77 := [mp~ #61 #80]: #58
+#89 := [mp #77 #88]: #86
+#155 := [mp #89 #154]: #150
+#187 := (not #150)
+#188 := (or #187 #184)
+#157 := (+ #22 #156)
+#158 := (<= #157 0::real)
+#159 := (not #158)
+#164 := (or #163 #159)
+#165 := (not #164)
+#189 := (or #187 #165)
+#191 := (iff #189 #188)
+#193 := (iff #188 #188)
+#194 := [rewrite]: #193
+#185 := (iff #165 #184)
+#182 := (iff #164 #181)
+#179 := (iff #159 #178)
+#176 := (iff #158 #173)
+#166 := (+ #156 #22)
+#169 := (<= #166 0::real)
+#174 := (iff #169 #173)
+#175 := [rewrite]: #174
+#170 := (iff #158 #169)
+#167 := (= #157 #166)
+#168 := [rewrite]: #167
+#171 := [monotonicity #168]: #170
+#177 := [trans #171 #175]: #176
+#180 := [monotonicity #177]: #179
+#183 := [monotonicity #180]: #182
+#186 := [monotonicity #183]: #185
+#192 := [monotonicity #186]: #191
+#195 := [trans #192 #194]: #191
+#190 := [quant-inst]: #189
+#196 := [mp #190 #195]: #188
+#220 := [unit-resolution #196 #155]: #184
+#197 := (or #181 #162)
+#198 := [def-axiom]: #197
+#221 := [unit-resolution #198 #220]: #162
+#223 := (or #206 #163 #222)
+#224 := [th-lemma]: #223
+#226 := [unit-resolution #224 #221]: #225
+#227 := [unit-resolution #226 #219]: #222
+#199 := (or #181 #173)
+#200 := [def-axiom]: #199
+#228 := [unit-resolution #200 #220]: #173
+[th-lemma #228 #227 #211]: false
+unsat
+06d4133eb0950a5f12d22480aa2707e31af2b064 86 0
#2 := false
#37 := 0::real
decl uf_2 :: (-> T2 T1 real)
@@ -3210,511 +3210,3 @@
#155 := [th-lemma]: #154
[unit-resolution #155 #60 #148 #144]: false
unsat
-V+IAyBZU/6QjYs6JkXx8LQ 57 0
-#2 := false
-#4 := 0::real
-decl uf_1 :: (-> T2 real)
-decl uf_2 :: (-> T1 T1 T2)
-decl uf_12 :: (-> T4 T1)
-decl uf_4 :: T4
-#11 := uf_4
-#39 := (uf_12 uf_4)
-decl uf_10 :: T4
-#27 := uf_10
-#38 := (uf_12 uf_10)
-#40 := (uf_2 #38 #39)
-#41 := (uf_1 #40)
-#264 := (>= #41 0::real)
-#266 := (not #264)
-#43 := (= #41 0::real)
-#44 := (not #43)
-#131 := [asserted]: #44
-#272 := (or #43 #266)
-#42 := (<= #41 0::real)
-#130 := [asserted]: #42
-#265 := (not #42)
-#270 := (or #43 #265 #266)
-#271 := [th-lemma]: #270
-#273 := [unit-resolution #271 #130]: #272
-#274 := [unit-resolution #273 #131]: #266
-#6 := (:var 0 T1)
-#5 := (:var 1 T1)
-#7 := (uf_2 #5 #6)
-#241 := (pattern #7)
-#8 := (uf_1 #7)
-#65 := (>= #8 0::real)
-#242 := (forall (vars (?x1 T1) (?x2 T1)) (:pat #241) #65)
-#66 := (forall (vars (?x1 T1) (?x2 T1)) #65)
-#245 := (iff #66 #242)
-#243 := (iff #65 #65)
-#244 := [refl]: #243
-#246 := [quant-intro #244]: #245
-#149 := (~ #66 #66)
-#151 := (~ #65 #65)
-#152 := [refl]: #151
-#150 := [nnf-pos #152]: #149
-#9 := (<= 0::real #8)
-#10 := (forall (vars (?x1 T1) (?x2 T1)) #9)
-#67 := (iff #10 #66)
-#63 := (iff #9 #65)
-#64 := [rewrite]: #63
-#68 := [quant-intro #64]: #67
-#60 := [asserted]: #10
-#69 := [mp #60 #68]: #66
-#147 := [mp~ #69 #150]: #66
-#247 := [mp #147 #246]: #242
-#267 := (not #242)
-#268 := (or #267 #264)
-#269 := [quant-inst]: #268
-[unit-resolution #269 #247 #274]: false
-unsat
-vqiyJ/qjGXZ3iOg6xftiug 15 0
-uf_1 -> val!0
-uf_2 -> val!1
-uf_3 -> val!2
-uf_5 -> val!15
-uf_6 -> val!26
-uf_4 -> {
- val!0 -> val!12
- val!1 -> val!13
- else -> val!13
-}
-uf_7 -> {
- val!6 -> val!31
- else -> val!31
-}
-sat
-mrZPJZyTokErrN6SYupisw 9 0
-uf_4 -> val!1
-uf_2 -> val!3
-uf_3 -> val!4
-uf_1 -> {
- val!5 -> val!6
- val!4 -> val!7
- else -> val!7
-}
-sat
-qkVrmXMcHAG5MLuJ9d8jXQ 211 0
-#2 := false
-#33 := 0::real
-decl uf_11 :: (-> T5 T6 real)
-decl uf_15 :: T6
-#28 := uf_15
-decl uf_16 :: T5
-#30 := uf_16
-#31 := (uf_11 uf_16 uf_15)
-decl uf_12 :: (-> T7 T8 T5)
-decl uf_14 :: T8
-#26 := uf_14
-decl uf_13 :: (-> T1 T7)
-decl uf_8 :: T1
-#16 := uf_8
-#25 := (uf_13 uf_8)
-#27 := (uf_12 #25 uf_14)
-#29 := (uf_11 #27 uf_15)
-#73 := -1::real
-#84 := (* -1::real #29)
-#85 := (+ #84 #31)
-#74 := (* -1::real #31)
-#75 := (+ #29 #74)
-#112 := (>= #75 0::real)
-#119 := (ite #112 #75 #85)
-#127 := (* -1::real #119)
-decl uf_17 :: T5
-#37 := uf_17
-#38 := (uf_11 uf_17 uf_15)
-#102 := -1/3::real
-#103 := (* -1/3::real #38)
-#128 := (+ #103 #127)
-#100 := 1/3::real
-#101 := (* 1/3::real #31)
-#129 := (+ #101 #128)
-#130 := (<= #129 0::real)
-#131 := (not #130)
-#40 := 3::real
-#39 := (- #31 #38)
-#41 := (/ #39 3::real)
-#32 := (- #29 #31)
-#35 := (- #32)
-#34 := (< #32 0::real)
-#36 := (ite #34 #35 #32)
-#42 := (< #36 #41)
-#136 := (iff #42 #131)
-#104 := (+ #101 #103)
-#78 := (< #75 0::real)
-#90 := (ite #78 #85 #75)
-#109 := (< #90 #104)
-#134 := (iff #109 #131)
-#124 := (< #119 #104)
-#132 := (iff #124 #131)
-#133 := [rewrite]: #132
-#125 := (iff #109 #124)
-#122 := (= #90 #119)
-#113 := (not #112)
-#116 := (ite #113 #85 #75)
-#120 := (= #116 #119)
-#121 := [rewrite]: #120
-#117 := (= #90 #116)
-#114 := (iff #78 #113)
-#115 := [rewrite]: #114
-#118 := [monotonicity #115]: #117
-#123 := [trans #118 #121]: #122
-#126 := [monotonicity #123]: #125
-#135 := [trans #126 #133]: #134
-#110 := (iff #42 #109)
-#107 := (= #41 #104)
-#93 := (* -1::real #38)
-#94 := (+ #31 #93)
-#97 := (/ #94 3::real)
-#105 := (= #97 #104)
-#106 := [rewrite]: #105
-#98 := (= #41 #97)
-#95 := (= #39 #94)
-#96 := [rewrite]: #95
-#99 := [monotonicity #96]: #98
-#108 := [trans #99 #106]: #107
-#91 := (= #36 #90)
-#76 := (= #32 #75)
-#77 := [rewrite]: #76
-#88 := (= #35 #85)
-#81 := (- #75)
-#86 := (= #81 #85)
-#87 := [rewrite]: #86
-#82 := (= #35 #81)
-#83 := [monotonicity #77]: #82
-#89 := [trans #83 #87]: #88
-#79 := (iff #34 #78)
-#80 := [monotonicity #77]: #79
-#92 := [monotonicity #80 #89 #77]: #91
-#111 := [monotonicity #92 #108]: #110
-#137 := [trans #111 #135]: #136
-#72 := [asserted]: #42
-#138 := [mp #72 #137]: #131
-decl uf_1 :: T1
-#4 := uf_1
-#43 := (uf_13 uf_1)
-#44 := (uf_12 #43 uf_14)
-#45 := (uf_11 #44 uf_15)
-#149 := (* -1::real #45)
-#150 := (+ #38 #149)
-#140 := (+ #93 #45)
-#161 := (<= #150 0::real)
-#168 := (ite #161 #140 #150)
-#176 := (* -1::real #168)
-#177 := (+ #103 #176)
-#178 := (+ #101 #177)
-#179 := (<= #178 0::real)
-#180 := (not #179)
-#46 := (- #45 #38)
-#48 := (- #46)
-#47 := (< #46 0::real)
-#49 := (ite #47 #48 #46)
-#50 := (< #49 #41)
-#185 := (iff #50 #180)
-#143 := (< #140 0::real)
-#155 := (ite #143 #150 #140)
-#158 := (< #155 #104)
-#183 := (iff #158 #180)
-#173 := (< #168 #104)
-#181 := (iff #173 #180)
-#182 := [rewrite]: #181
-#174 := (iff #158 #173)
-#171 := (= #155 #168)
-#162 := (not #161)
-#165 := (ite #162 #150 #140)
-#169 := (= #165 #168)
-#170 := [rewrite]: #169
-#166 := (= #155 #165)
-#163 := (iff #143 #162)
-#164 := [rewrite]: #163
-#167 := [monotonicity #164]: #166
-#172 := [trans #167 #170]: #171
-#175 := [monotonicity #172]: #174
-#184 := [trans #175 #182]: #183
-#159 := (iff #50 #158)
-#156 := (= #49 #155)
-#141 := (= #46 #140)
-#142 := [rewrite]: #141
-#153 := (= #48 #150)
-#146 := (- #140)
-#151 := (= #146 #150)
-#152 := [rewrite]: #151
-#147 := (= #48 #146)
-#148 := [monotonicity #142]: #147
-#154 := [trans #148 #152]: #153
-#144 := (iff #47 #143)
-#145 := [monotonicity #142]: #144
-#157 := [monotonicity #145 #154 #142]: #156
-#160 := [monotonicity #157 #108]: #159
-#186 := [trans #160 #184]: #185
-#139 := [asserted]: #50
-#187 := [mp #139 #186]: #180
-#299 := (+ #140 #176)
-#300 := (<= #299 0::real)
-#290 := (= #140 #168)
-#329 := [hypothesis]: #162
-#191 := (+ #29 #149)
-#192 := (<= #191 0::real)
-#51 := (<= #29 #45)
-#193 := (iff #51 #192)
-#194 := [rewrite]: #193
-#188 := [asserted]: #51
-#195 := [mp #188 #194]: #192
-#298 := (+ #75 #127)
-#301 := (<= #298 0::real)
-#284 := (= #75 #119)
-#302 := [hypothesis]: #113
-#296 := (+ #85 #127)
-#297 := (<= #296 0::real)
-#285 := (= #85 #119)
-#288 := (or #112 #285)
-#289 := [def-axiom]: #288
-#303 := [unit-resolution #289 #302]: #285
-#304 := (not #285)
-#305 := (or #304 #297)
-#306 := [th-lemma]: #305
-#307 := [unit-resolution #306 #303]: #297
-#315 := (not #290)
-#310 := (not #300)
-#311 := (or #310 #112)
-#308 := [hypothesis]: #300
-#309 := [th-lemma #308 #307 #138 #302 #187 #195]: false
-#312 := [lemma #309]: #311
-#322 := [unit-resolution #312 #302]: #310
-#316 := (or #315 #300)
-#313 := [hypothesis]: #310
-#314 := [hypothesis]: #290
-#317 := [th-lemma]: #316
-#318 := [unit-resolution #317 #314 #313]: false
-#319 := [lemma #318]: #316
-#323 := [unit-resolution #319 #322]: #315
-#292 := (or #162 #290)
-#293 := [def-axiom]: #292
-#324 := [unit-resolution #293 #323]: #162
-#325 := [th-lemma #324 #307 #138 #302 #195]: false
-#326 := [lemma #325]: #112
-#286 := (or #113 #284)
-#287 := [def-axiom]: #286
-#330 := [unit-resolution #287 #326]: #284
-#331 := (not #284)
-#332 := (or #331 #301)
-#333 := [th-lemma]: #332
-#334 := [unit-resolution #333 #330]: #301
-#335 := [th-lemma #326 #334 #195 #329 #138]: false
-#336 := [lemma #335]: #161
-#327 := [unit-resolution #293 #336]: #290
-#328 := [unit-resolution #319 #327]: #300
-[th-lemma #326 #334 #195 #328 #187 #138]: false
-unsat
-WrIjxhfF5EcRCmS6xKG3XQ 211 0
-#2 := false
-#33 := 0::real
-decl uf_11 :: (-> T5 T6 real)
-decl uf_15 :: T6
-#28 := uf_15
-decl uf_16 :: T5
-#30 := uf_16
-#31 := (uf_11 uf_16 uf_15)
-decl uf_12 :: (-> T7 T8 T5)
-decl uf_14 :: T8
-#26 := uf_14
-decl uf_13 :: (-> T1 T7)
-decl uf_8 :: T1
-#16 := uf_8
-#25 := (uf_13 uf_8)
-#27 := (uf_12 #25 uf_14)
-#29 := (uf_11 #27 uf_15)
-#73 := -1::real
-#84 := (* -1::real #29)
-#85 := (+ #84 #31)
-#74 := (* -1::real #31)
-#75 := (+ #29 #74)
-#112 := (>= #75 0::real)
-#119 := (ite #112 #75 #85)
-#127 := (* -1::real #119)
-decl uf_17 :: T5
-#37 := uf_17
-#38 := (uf_11 uf_17 uf_15)
-#102 := -1/3::real
-#103 := (* -1/3::real #38)
-#128 := (+ #103 #127)
-#100 := 1/3::real
-#101 := (* 1/3::real #31)
-#129 := (+ #101 #128)
-#130 := (<= #129 0::real)
-#131 := (not #130)
-#40 := 3::real
-#39 := (- #31 #38)
-#41 := (/ #39 3::real)
-#32 := (- #29 #31)
-#35 := (- #32)
-#34 := (< #32 0::real)
-#36 := (ite #34 #35 #32)
-#42 := (< #36 #41)
-#136 := (iff #42 #131)
-#104 := (+ #101 #103)
-#78 := (< #75 0::real)
-#90 := (ite #78 #85 #75)
-#109 := (< #90 #104)
-#134 := (iff #109 #131)
-#124 := (< #119 #104)
-#132 := (iff #124 #131)
-#133 := [rewrite]: #132
-#125 := (iff #109 #124)
-#122 := (= #90 #119)
-#113 := (not #112)
-#116 := (ite #113 #85 #75)
-#120 := (= #116 #119)
-#121 := [rewrite]: #120
-#117 := (= #90 #116)
-#114 := (iff #78 #113)
-#115 := [rewrite]: #114
-#118 := [monotonicity #115]: #117
-#123 := [trans #118 #121]: #122
-#126 := [monotonicity #123]: #125
-#135 := [trans #126 #133]: #134
-#110 := (iff #42 #109)
-#107 := (= #41 #104)
-#93 := (* -1::real #38)
-#94 := (+ #31 #93)
-#97 := (/ #94 3::real)
-#105 := (= #97 #104)
-#106 := [rewrite]: #105
-#98 := (= #41 #97)
-#95 := (= #39 #94)
-#96 := [rewrite]: #95
-#99 := [monotonicity #96]: #98
-#108 := [trans #99 #106]: #107
-#91 := (= #36 #90)
-#76 := (= #32 #75)
-#77 := [rewrite]: #76
-#88 := (= #35 #85)
-#81 := (- #75)
-#86 := (= #81 #85)
-#87 := [rewrite]: #86
-#82 := (= #35 #81)
-#83 := [monotonicity #77]: #82
-#89 := [trans #83 #87]: #88
-#79 := (iff #34 #78)
-#80 := [monotonicity #77]: #79
-#92 := [monotonicity #80 #89 #77]: #91
-#111 := [monotonicity #92 #108]: #110
-#137 := [trans #111 #135]: #136
-#72 := [asserted]: #42
-#138 := [mp #72 #137]: #131
-decl uf_1 :: T1
-#4 := uf_1
-#43 := (uf_13 uf_1)
-#44 := (uf_12 #43 uf_14)
-#45 := (uf_11 #44 uf_15)
-#149 := (* -1::real #45)
-#150 := (+ #38 #149)
-#140 := (+ #93 #45)
-#161 := (<= #150 0::real)
-#168 := (ite #161 #140 #150)
-#176 := (* -1::real #168)
-#177 := (+ #103 #176)
-#178 := (+ #101 #177)
-#179 := (<= #178 0::real)
-#180 := (not #179)
-#46 := (- #45 #38)
-#48 := (- #46)
-#47 := (< #46 0::real)
-#49 := (ite #47 #48 #46)
-#50 := (< #49 #41)
-#185 := (iff #50 #180)
-#143 := (< #140 0::real)
-#155 := (ite #143 #150 #140)
-#158 := (< #155 #104)
-#183 := (iff #158 #180)
-#173 := (< #168 #104)
-#181 := (iff #173 #180)
-#182 := [rewrite]: #181
-#174 := (iff #158 #173)
-#171 := (= #155 #168)
-#162 := (not #161)
-#165 := (ite #162 #150 #140)
-#169 := (= #165 #168)
-#170 := [rewrite]: #169
-#166 := (= #155 #165)
-#163 := (iff #143 #162)
-#164 := [rewrite]: #163
-#167 := [monotonicity #164]: #166
-#172 := [trans #167 #170]: #171
-#175 := [monotonicity #172]: #174
-#184 := [trans #175 #182]: #183
-#159 := (iff #50 #158)
-#156 := (= #49 #155)
-#141 := (= #46 #140)
-#142 := [rewrite]: #141
-#153 := (= #48 #150)
-#146 := (- #140)
-#151 := (= #146 #150)
-#152 := [rewrite]: #151
-#147 := (= #48 #146)
-#148 := [monotonicity #142]: #147
-#154 := [trans #148 #152]: #153
-#144 := (iff #47 #143)
-#145 := [monotonicity #142]: #144
-#157 := [monotonicity #145 #154 #142]: #156
-#160 := [monotonicity #157 #108]: #159
-#186 := [trans #160 #184]: #185
-#139 := [asserted]: #50
-#187 := [mp #139 #186]: #180
-#299 := (+ #140 #176)
-#300 := (<= #299 0::real)
-#290 := (= #140 #168)
-#329 := [hypothesis]: #162
-#191 := (+ #29 #149)
-#192 := (<= #191 0::real)
-#51 := (<= #29 #45)
-#193 := (iff #51 #192)
-#194 := [rewrite]: #193
-#188 := [asserted]: #51
-#195 := [mp #188 #194]: #192
-#298 := (+ #75 #127)
-#301 := (<= #298 0::real)
-#284 := (= #75 #119)
-#302 := [hypothesis]: #113
-#296 := (+ #85 #127)
-#297 := (<= #296 0::real)
-#285 := (= #85 #119)
-#288 := (or #112 #285)
-#289 := [def-axiom]: #288
-#303 := [unit-resolution #289 #302]: #285
-#304 := (not #285)
-#305 := (or #304 #297)
-#306 := [th-lemma]: #305
-#307 := [unit-resolution #306 #303]: #297
-#315 := (not #290)
-#310 := (not #300)
-#311 := (or #310 #112)
-#308 := [hypothesis]: #300
-#309 := [th-lemma #308 #307 #138 #302 #187 #195]: false
-#312 := [lemma #309]: #311
-#322 := [unit-resolution #312 #302]: #310
-#316 := (or #315 #300)
-#313 := [hypothesis]: #310
-#314 := [hypothesis]: #290
-#317 := [th-lemma]: #316
-#318 := [unit-resolution #317 #314 #313]: false
-#319 := [lemma #318]: #316
-#323 := [unit-resolution #319 #322]: #315
-#292 := (or #162 #290)
-#293 := [def-axiom]: #292
-#324 := [unit-resolution #293 #323]: #162
-#325 := [th-lemma #324 #307 #138 #302 #195]: false
-#326 := [lemma #325]: #112
-#286 := (or #113 #284)
-#287 := [def-axiom]: #286
-#330 := [unit-resolution #287 #326]: #284
-#331 := (not #284)
-#332 := (or #331 #301)
-#333 := [th-lemma]: #332
-#334 := [unit-resolution #333 #330]: #301
-#335 := [th-lemma #326 #334 #195 #329 #138]: false
-#336 := [lemma #335]: #161
-#327 := [unit-resolution #293 #336]: #290
-#328 := [unit-resolution #319 #327]: #300
-[th-lemma #326 #334 #195 #328 #187 #138]: false
-unsat
--- a/src/HOL/Multivariate_Analysis/Integration.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOL/Multivariate_Analysis/Integration.thy Thu Mar 25 17:56:31 2010 +0100
@@ -8,7 +8,7 @@
begin
declare [[smt_certificates="~~/src/HOL/Multivariate_Analysis/Integration.cert"]]
-declare [[smt_record=true]]
+declare [[smt_record=false]]
declare [[z3_proofs=true]]
lemma conjunctD2: assumes "a \<and> b" shows a b using assms by auto
--- a/src/HOL/Predicate_Compile.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOL/Predicate_Compile.thy Thu Mar 25 17:56:31 2010 +0100
@@ -1,4 +1,3 @@
-
(* Title: HOL/Predicate_Compile.thy
Author: Stefan Berghofer, Lukas Bulwahn, Florian Haftmann, TU Muenchen
*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Examples.thy Thu Mar 25 17:56:31 2010 +0100
@@ -0,0 +1,1278 @@
+theory Predicate_Compile_Examples
+imports Predicate_Compile_Alternative_Defs
+begin
+
+subsection {* Basic predicates *}
+
+inductive False' :: "bool"
+
+code_pred (expected_modes: bool) False' .
+code_pred [dseq] False' .
+code_pred [random_dseq] False' .
+
+values [expected "{}" pred] "{x. False'}"
+values [expected "{}" dseq 1] "{x. False'}"
+values [expected "{}" random_dseq 1, 1, 1] "{x. False'}"
+
+value "False'"
+
+
+inductive True' :: "bool"
+where
+ "True ==> True'"
+
+code_pred True' .
+code_pred [dseq] True' .
+code_pred [random_dseq] True' .
+
+thm True'.equation
+thm True'.dseq_equation
+thm True'.random_dseq_equation
+values [expected "{()}" ]"{x. True'}"
+values [expected "{}" dseq 0] "{x. True'}"
+values [expected "{()}" dseq 1] "{x. True'}"
+values [expected "{()}" dseq 2] "{x. True'}"
+values [expected "{}" random_dseq 1, 1, 0] "{x. True'}"
+values [expected "{}" random_dseq 1, 1, 1] "{x. True'}"
+values [expected "{()}" random_dseq 1, 1, 2] "{x. True'}"
+values [expected "{()}" random_dseq 1, 1, 3] "{x. True'}"
+
+inductive EmptySet :: "'a \<Rightarrow> bool"
+
+code_pred (expected_modes: o => bool, i => bool) EmptySet .
+
+definition EmptySet' :: "'a \<Rightarrow> bool"
+where "EmptySet' = {}"
+
+code_pred (expected_modes: o => bool, i => bool) [inductify] EmptySet' .
+
+inductive EmptyRel :: "'a \<Rightarrow> 'b \<Rightarrow> bool"
+
+code_pred (expected_modes: o => o => bool, i => o => bool, o => i => bool, i => i => bool) EmptyRel .
+
+inductive EmptyClosure :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
+for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+
+code_pred
+ (expected_modes: (o => o => bool) => o => o => bool, (o => o => bool) => i => o => bool,
+ (o => o => bool) => o => i => bool, (o => o => bool) => i => i => bool,
+ (i => o => bool) => o => o => bool, (i => o => bool) => i => o => bool,
+ (i => o => bool) => o => i => bool, (i => o => bool) => i => i => bool,
+ (o => i => bool) => o => o => bool, (o => i => bool) => i => o => bool,
+ (o => i => bool) => o => i => bool, (o => i => bool) => i => i => bool,
+ (i => i => bool) => o => o => bool, (i => i => bool) => i => o => bool,
+ (i => i => bool) => o => i => bool, (i => i => bool) => i => i => bool)
+ EmptyClosure .
+
+thm EmptyClosure.equation
+
+(* TODO: inductive package is broken!
+inductive False'' :: "bool"
+where
+ "False \<Longrightarrow> False''"
+
+code_pred (expected_modes: []) False'' .
+
+inductive EmptySet'' :: "'a \<Rightarrow> bool"
+where
+ "False \<Longrightarrow> EmptySet'' x"
+
+code_pred (expected_modes: [1]) EmptySet'' .
+code_pred (expected_modes: [], [1]) [inductify] EmptySet'' .
+*)
+
+consts a' :: 'a
+
+inductive Fact :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+where
+"Fact a' a'"
+
+code_pred (expected_modes: o => o => bool, i => o => bool, o => i => bool, i => i => bool) Fact .
+
+inductive zerozero :: "nat * nat => bool"
+where
+ "zerozero (0, 0)"
+
+code_pred (expected_modes: i => bool, i * o => bool, o * i => bool, o => bool) zerozero .
+code_pred [dseq] zerozero .
+code_pred [random_dseq] zerozero .
+
+thm zerozero.equation
+thm zerozero.dseq_equation
+thm zerozero.random_dseq_equation
+
+text {* We expect the user to expand the tuples in the values command.
+The following values command is not supported. *}
+(*values "{x. zerozero x}" *)
+text {* Instead, the user must type *}
+values "{(x, y). zerozero (x, y)}"
+
+values [expected "{}" dseq 0] "{(x, y). zerozero (x, y)}"
+values [expected "{(0::nat, 0::nat)}" dseq 1] "{(x, y). zerozero (x, y)}"
+values [expected "{(0::nat, 0::nat)}" dseq 2] "{(x, y). zerozero (x, y)}"
+values [expected "{}" random_dseq 1, 1, 2] "{(x, y). zerozero (x, y)}"
+values [expected "{(0::nat, 0:: nat)}" random_dseq 1, 1, 3] "{(x, y). zerozero (x, y)}"
+
+inductive nested_tuples :: "((int * int) * int * int) => bool"
+where
+ "nested_tuples ((0, 1), 2, 3)"
+
+code_pred nested_tuples .
+
+inductive JamesBond :: "nat => int => code_numeral => bool"
+where
+ "JamesBond 0 0 7"
+
+code_pred JamesBond .
+
+values [expected "{(0::nat, 0::int , 7::code_numeral)}"] "{(a, b, c). JamesBond a b c}"
+values [expected "{(0::nat, 7::code_numeral, 0:: int)}"] "{(a, c, b). JamesBond a b c}"
+values [expected "{(0::int, 0::nat, 7::code_numeral)}"] "{(b, a, c). JamesBond a b c}"
+values [expected "{(0::int, 7::code_numeral, 0::nat)}"] "{(b, c, a). JamesBond a b c}"
+values [expected "{(7::code_numeral, 0::nat, 0::int)}"] "{(c, a, b). JamesBond a b c}"
+values [expected "{(7::code_numeral, 0::int, 0::nat)}"] "{(c, b, a). JamesBond a b c}"
+
+values [expected "{(7::code_numeral, 0::int)}"] "{(a, b). JamesBond 0 b a}"
+values [expected "{(7::code_numeral, 0::nat)}"] "{(c, a). JamesBond a 0 c}"
+values [expected "{(0::nat, 7::code_numeral)}"] "{(a, c). JamesBond a 0 c}"
+
+
+subsection {* Alternative Rules *}
+
+datatype char = C | D | E | F | G | H
+
+inductive is_C_or_D
+where
+ "(x = C) \<or> (x = D) ==> is_C_or_D x"
+
+code_pred (expected_modes: i => bool) is_C_or_D .
+thm is_C_or_D.equation
+
+inductive is_D_or_E
+where
+ "(x = D) \<or> (x = E) ==> is_D_or_E x"
+
+lemma [code_pred_intro]:
+ "is_D_or_E D"
+by (auto intro: is_D_or_E.intros)
+
+lemma [code_pred_intro]:
+ "is_D_or_E E"
+by (auto intro: is_D_or_E.intros)
+
+code_pred (expected_modes: o => bool, i => bool) is_D_or_E
+proof -
+ case is_D_or_E
+ from this(1) show thesis
+ proof
+ fix xa
+ assume x: "x = xa"
+ assume "xa = D \<or> xa = E"
+ from this show thesis
+ proof
+ assume "xa = D" from this x is_D_or_E(2) show thesis by simp
+ next
+ assume "xa = E" from this x is_D_or_E(3) show thesis by simp
+ qed
+ qed
+qed
+
+thm is_D_or_E.equation
+
+inductive is_F_or_G
+where
+ "x = F \<or> x = G ==> is_F_or_G x"
+
+lemma [code_pred_intro]:
+ "is_F_or_G F"
+by (auto intro: is_F_or_G.intros)
+
+lemma [code_pred_intro]:
+ "is_F_or_G G"
+by (auto intro: is_F_or_G.intros)
+
+inductive is_FGH
+where
+ "is_F_or_G x ==> is_FGH x"
+| "is_FGH H"
+
+text {* Compilation of is_FGH requires elimination rule for is_F_or_G *}
+
+code_pred (expected_modes: o => bool, i => bool) is_FGH
+proof -
+ case is_F_or_G
+ from this(1) show thesis
+ proof
+ fix xa
+ assume x: "x = xa"
+ assume "xa = F \<or> xa = G"
+ from this show thesis
+ proof
+ assume "xa = F"
+ from this x is_F_or_G(2) show thesis by simp
+ next
+ assume "xa = G"
+ from this x is_F_or_G(3) show thesis by simp
+ qed
+ qed
+qed
+
+subsection {* Preprocessor Inlining *}
+
+definition "equals == (op =)"
+
+inductive zerozero' :: "nat * nat => bool" where
+ "equals (x, y) (0, 0) ==> zerozero' (x, y)"
+
+code_pred (expected_modes: i => bool) zerozero' .
+
+lemma zerozero'_eq: "zerozero' x == zerozero x"
+proof -
+ have "zerozero' = zerozero"
+ apply (auto simp add: mem_def)
+ apply (cases rule: zerozero'.cases)
+ apply (auto simp add: equals_def intro: zerozero.intros)
+ apply (cases rule: zerozero.cases)
+ apply (auto simp add: equals_def intro: zerozero'.intros)
+ done
+ from this show "zerozero' x == zerozero x" by auto
+qed
+
+declare zerozero'_eq [code_pred_inline]
+
+definition "zerozero'' x == zerozero' x"
+
+text {* if preprocessing fails, zerozero'' will not have all modes. *}
+
+code_pred (expected_modes: i * i => bool, i * o => bool, o * i => bool, o => bool) [inductify] zerozero'' .
+
+subsection {* Sets and Numerals *}
+
+definition
+ "one_or_two = {Suc 0, (Suc (Suc 0))}"
+
+code_pred [inductify] one_or_two .
+
+code_pred [dseq] one_or_two .
+code_pred [random_dseq] one_or_two .
+thm one_or_two.dseq_equation
+values [expected "{Suc 0::nat, 2::nat}"] "{x. one_or_two x}"
+values [random_dseq 0,0,10] 3 "{x. one_or_two x}"
+
+inductive one_or_two' :: "nat => bool"
+where
+ "one_or_two' 1"
+| "one_or_two' 2"
+
+code_pred one_or_two' .
+thm one_or_two'.equation
+
+values "{x. one_or_two' x}"
+
+definition one_or_two'':
+ "one_or_two'' == {1, (2::nat)}"
+
+code_pred [inductify] one_or_two'' .
+thm one_or_two''.equation
+
+values "{x. one_or_two'' x}"
+
+subsection {* even predicate *}
+
+inductive even :: "nat \<Rightarrow> bool" and odd :: "nat \<Rightarrow> bool" where
+ "even 0"
+ | "even n \<Longrightarrow> odd (Suc n)"
+ | "odd n \<Longrightarrow> even (Suc n)"
+
+code_pred (expected_modes: i => bool, o => bool) even .
+code_pred [dseq] even .
+code_pred [random_dseq] even .
+
+thm odd.equation
+thm even.equation
+thm odd.dseq_equation
+thm even.dseq_equation
+thm odd.random_dseq_equation
+thm even.random_dseq_equation
+
+values "{x. even 2}"
+values "{x. odd 2}"
+values 10 "{n. even n}"
+values 10 "{n. odd n}"
+values [expected "{}" dseq 2] "{x. even 6}"
+values [expected "{}" dseq 6] "{x. even 6}"
+values [expected "{()}" dseq 7] "{x. even 6}"
+values [dseq 2] "{x. odd 7}"
+values [dseq 6] "{x. odd 7}"
+values [dseq 7] "{x. odd 7}"
+values [expected "{()}" dseq 8] "{x. odd 7}"
+
+values [expected "{}" dseq 0] 8 "{x. even x}"
+values [expected "{0::nat}" dseq 1] 8 "{x. even x}"
+values [expected "{0::nat, 2}" dseq 3] 8 "{x. even x}"
+values [expected "{0::nat, 2}" dseq 4] 8 "{x. even x}"
+values [expected "{0::nat, 2, 4}" dseq 6] 8 "{x. even x}"
+
+values [random_dseq 1, 1, 0] 8 "{x. even x}"
+values [random_dseq 1, 1, 1] 8 "{x. even x}"
+values [random_dseq 1, 1, 2] 8 "{x. even x}"
+values [random_dseq 1, 1, 3] 8 "{x. even x}"
+values [random_dseq 1, 1, 6] 8 "{x. even x}"
+
+values [expected "{}" random_dseq 1, 1, 7] "{x. odd 7}"
+values [random_dseq 1, 1, 8] "{x. odd 7}"
+values [random_dseq 1, 1, 9] "{x. odd 7}"
+
+definition odd' where "odd' x == \<not> even x"
+
+code_pred (expected_modes: i => bool) [inductify] odd' .
+code_pred [dseq inductify] odd' .
+code_pred [random_dseq inductify] odd' .
+
+values [expected "{}" dseq 2] "{x. odd' 7}"
+values [expected "{()}" dseq 9] "{x. odd' 7}"
+values [expected "{}" dseq 2] "{x. odd' 8}"
+values [expected "{}" dseq 10] "{x. odd' 8}"
+
+
+inductive is_even :: "nat \<Rightarrow> bool"
+where
+ "n mod 2 = 0 \<Longrightarrow> is_even n"
+
+code_pred (expected_modes: i => bool) is_even .
+
+subsection {* append predicate *}
+
+inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
+ "append [] xs xs"
+ | "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)"
+
+code_pred (modes: i => i => o => bool as "concat", o => o => i => bool as "slice", o => i => i => bool as prefix,
+ i => o => i => bool as suffix, i => i => i => bool) append .
+code_pred [dseq] append .
+code_pred [random_dseq] append .
+
+thm append.equation
+thm append.dseq_equation
+thm append.random_dseq_equation
+
+values "{(ys, xs). append xs ys [0, Suc 0, 2]}"
+values "{zs. append [0, Suc 0, 2] [17, 8] zs}"
+values "{ys. append [0, Suc 0, 2] ys [0, Suc 0, 2, 17, 0, 5]}"
+
+values [expected "{}" dseq 0] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
+values [expected "{(([]::nat list), [Suc 0, 2, 3, 4, (5::nat)])}" dseq 1] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
+values [dseq 4] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
+values [dseq 6] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
+values [random_dseq 1, 1, 4] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
+values [random_dseq 1, 1, 1] 10 "{(xs, ys, zs::int list). append xs ys zs}"
+values [random_dseq 1, 1, 3] 10 "{(xs, ys, zs::int list). append xs ys zs}"
+values [random_dseq 3, 1, 3] 10 "{(xs, ys, zs::int list). append xs ys zs}"
+values [random_dseq 1, 3, 3] 10 "{(xs, ys, zs::int list). append xs ys zs}"
+values [random_dseq 1, 1, 4] 10 "{(xs, ys, zs::int list). append xs ys zs}"
+
+value [code] "Predicate.the (concat [0::int, 1, 2] [3, 4, 5])"
+value [code] "Predicate.the (slice ([]::int list))"
+
+
+text {* tricky case with alternative rules *}
+
+inductive append2
+where
+ "append2 [] xs xs"
+| "append2 xs ys zs \<Longrightarrow> append2 (x # xs) ys (x # zs)"
+
+lemma append2_Nil: "append2 [] (xs::'b list) xs"
+ by (simp add: append2.intros(1))
+
+lemmas [code_pred_intro] = append2_Nil append2.intros(2)
+
+code_pred (expected_modes: i => i => o => bool, o => o => i => bool, o => i => i => bool,
+ i => o => i => bool, i => i => i => bool) append2
+proof -
+ case append2
+ from append2(1) show thesis
+ proof
+ fix xs
+ assume "xa = []" "xb = xs" "xc = xs"
+ from this append2(2) show thesis by simp
+ next
+ fix xs ys zs x
+ assume "xa = x # xs" "xb = ys" "xc = x # zs" "append2 xs ys zs"
+ from this append2(3) show thesis by fastsimp
+ qed
+qed
+
+inductive tupled_append :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
+where
+ "tupled_append ([], xs, xs)"
+| "tupled_append (xs, ys, zs) \<Longrightarrow> tupled_append (x # xs, ys, x # zs)"
+
+code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,
+ i * o * i => bool, i * i * i => bool) tupled_append .
+code_pred [random_dseq] tupled_append .
+thm tupled_append.equation
+
+values "{xs. tupled_append ([(1::nat), 2, 3], [4, 5], xs)}"
+
+inductive tupled_append'
+where
+"tupled_append' ([], xs, xs)"
+| "[| ys = fst (xa, y); x # zs = snd (xa, y);
+ tupled_append' (xs, ys, zs) |] ==> tupled_append' (x # xs, xa, y)"
+
+code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,
+ i * o * i => bool, i * i * i => bool) tupled_append' .
+thm tupled_append'.equation
+
+inductive tupled_append'' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
+where
+ "tupled_append'' ([], xs, xs)"
+| "ys = fst yszs ==> x # zs = snd yszs ==> tupled_append'' (xs, ys, zs) \<Longrightarrow> tupled_append'' (x # xs, yszs)"
+
+code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,
+ i * o * i => bool, i * i * i => bool) tupled_append'' .
+thm tupled_append''.equation
+
+inductive tupled_append''' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
+where
+ "tupled_append''' ([], xs, xs)"
+| "yszs = (ys, zs) ==> tupled_append''' (xs, yszs) \<Longrightarrow> tupled_append''' (x # xs, ys, x # zs)"
+
+code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,
+ i * o * i => bool, i * i * i => bool) tupled_append''' .
+thm tupled_append'''.equation
+
+subsection {* map_ofP predicate *}
+
+inductive map_ofP :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
+where
+ "map_ofP ((a, b)#xs) a b"
+| "map_ofP xs a b \<Longrightarrow> map_ofP (x#xs) a b"
+
+code_pred (expected_modes: i => o => o => bool, i => i => o => bool, i => o => i => bool, i => i => i => bool) map_ofP .
+thm map_ofP.equation
+
+subsection {* filter predicate *}
+
+inductive filter1
+for P
+where
+ "filter1 P [] []"
+| "P x ==> filter1 P xs ys ==> filter1 P (x#xs) (x#ys)"
+| "\<not> P x ==> filter1 P xs ys ==> filter1 P (x#xs) ys"
+
+code_pred (expected_modes: (i => bool) => i => o => bool, (i => bool) => i => i => bool) filter1 .
+code_pred [dseq] filter1 .
+code_pred [random_dseq] filter1 .
+
+thm filter1.equation
+
+values [expected "{[0::nat, 2, 4]}"] "{xs. filter1 even [0, 1, 2, 3, 4] xs}"
+values [expected "{}" dseq 9] "{xs. filter1 even [0, 1, 2, 3, 4] xs}"
+values [expected "{[0::nat, 2, 4]}" dseq 10] "{xs. filter1 even [0, 1, 2, 3, 4] xs}"
+
+inductive filter2
+where
+ "filter2 P [] []"
+| "P x ==> filter2 P xs ys ==> filter2 P (x#xs) (x#ys)"
+| "\<not> P x ==> filter2 P xs ys ==> filter2 P (x#xs) ys"
+
+code_pred (expected_modes: (i => bool) => i => i => bool, (i => bool) => i => o => bool) filter2 .
+code_pred [dseq] filter2 .
+code_pred [random_dseq] filter2 .
+
+thm filter2.equation
+thm filter2.random_dseq_equation
+
+(*
+inductive filter3
+for P
+where
+ "List.filter P xs = ys ==> filter3 P xs ys"
+
+code_pred (expected_modes: (o => bool) => i => o => bool, (o => bool) => i => i => bool , (i => bool) => i => o => bool, (i => bool) => i => i => bool) [skip_proof] filter3 .
+
+code_pred [dseq] filter3 .
+thm filter3.dseq_equation
+*)
+(*
+inductive filter4
+where
+ "List.filter P xs = ys ==> filter4 P xs ys"
+
+code_pred (expected_modes: i => i => o => bool, i => i => i => bool) filter4 .
+(*code_pred [depth_limited] filter4 .*)
+(*code_pred [random] filter4 .*)
+*)
+subsection {* reverse predicate *}
+
+inductive rev where
+ "rev [] []"
+ | "rev xs xs' ==> append xs' [x] ys ==> rev (x#xs) ys"
+
+code_pred (expected_modes: i => o => bool, o => i => bool, i => i => bool) rev .
+
+thm rev.equation
+
+values "{xs. rev [0, 1, 2, 3::nat] xs}"
+
+inductive tupled_rev where
+ "tupled_rev ([], [])"
+| "tupled_rev (xs, xs') \<Longrightarrow> tupled_append (xs', [x], ys) \<Longrightarrow> tupled_rev (x#xs, ys)"
+
+code_pred (expected_modes: i * o => bool, o * i => bool, i * i => bool) tupled_rev .
+thm tupled_rev.equation
+
+subsection {* partition predicate *}
+
+inductive partition :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
+ for f where
+ "partition f [] [] []"
+ | "f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) (x # ys) zs"
+ | "\<not> f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) ys (x # zs)"
+
+code_pred (expected_modes: (i => bool) => i => o => o => bool, (i => bool) => o => i => i => bool,
+ (i => bool) => i => i => o => bool, (i => bool) => i => o => i => bool, (i => bool) => i => i => i => bool)
+ partition .
+code_pred [dseq] partition .
+code_pred [random_dseq] partition .
+
+values 10 "{(ys, zs). partition is_even
+ [0, Suc 0, 2, 3, 4, 5, 6, 7] ys zs}"
+values 10 "{zs. partition is_even zs [0, 2] [3, 5]}"
+values 10 "{zs. partition is_even zs [0, 7] [3, 5]}"
+
+inductive tupled_partition :: "('a \<Rightarrow> bool) \<Rightarrow> ('a list \<times> 'a list \<times> 'a list) \<Rightarrow> bool"
+ for f where
+ "tupled_partition f ([], [], [])"
+ | "f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, x # ys, zs)"
+ | "\<not> f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, ys, x # zs)"
+
+code_pred (expected_modes: (i => bool) => i => bool, (i => bool) => (i * i * o) => bool, (i => bool) => (i * o * i) => bool,
+ (i => bool) => (o * i * i) => bool, (i => bool) => (i * o * o) => bool) tupled_partition .
+
+thm tupled_partition.equation
+
+lemma [code_pred_intro]:
+ "r a b \<Longrightarrow> tranclp r a b"
+ "r a b \<Longrightarrow> tranclp r b c \<Longrightarrow> tranclp r a c"
+ by auto
+
+subsection {* transitive predicate *}
+
+text {* Also look at the tabled transitive closure in the Library *}
+
+code_pred (modes: (i => o => bool) => i => i => bool, (i => o => bool) => i => o => bool as forwards_trancl,
+ (o => i => bool) => i => i => bool, (o => i => bool) => o => i => bool as backwards_trancl, (o => o => bool) => i => i => bool, (o => o => bool) => i => o => bool,
+ (o => o => bool) => o => i => bool, (o => o => bool) => o => o => bool) tranclp
+proof -
+ case tranclp
+ from this converse_tranclpE[OF this(1)] show thesis by metis
+qed
+
+
+code_pred [dseq] tranclp .
+code_pred [random_dseq] tranclp .
+thm tranclp.equation
+thm tranclp.random_dseq_equation
+
+inductive rtrancl' :: "'a => 'a => ('a => 'a => bool) => bool"
+where
+ "rtrancl' x x r"
+| "r x y ==> rtrancl' y z r ==> rtrancl' x z r"
+
+code_pred [random_dseq] rtrancl' .
+
+thm rtrancl'.random_dseq_equation
+
+inductive rtrancl'' :: "('a * 'a * ('a \<Rightarrow> 'a \<Rightarrow> bool)) \<Rightarrow> bool"
+where
+ "rtrancl'' (x, x, r)"
+| "r x y \<Longrightarrow> rtrancl'' (y, z, r) \<Longrightarrow> rtrancl'' (x, z, r)"
+
+code_pred rtrancl'' .
+
+inductive rtrancl''' :: "('a * ('a * 'a) * ('a * 'a => bool)) => bool"
+where
+ "rtrancl''' (x, (x, x), r)"
+| "r (x, y) ==> rtrancl''' (y, (z, z), r) ==> rtrancl''' (x, (z, z), r)"
+
+code_pred rtrancl''' .
+
+
+inductive succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where
+ "succ 0 1"
+ | "succ m n \<Longrightarrow> succ (Suc m) (Suc n)"
+
+code_pred (modes: i => i => bool, i => o => bool, o => i => bool, o => o => bool) succ .
+code_pred [random_dseq] succ .
+thm succ.equation
+thm succ.random_dseq_equation
+
+values 10 "{(m, n). succ n m}"
+values "{m. succ 0 m}"
+values "{m. succ m 0}"
+
+text {* values command needs mode annotation of the parameter succ
+to disambiguate which mode is to be chosen. *}
+
+values [mode: i => o => bool] 20 "{n. tranclp succ 10 n}"
+values [mode: o => i => bool] 10 "{n. tranclp succ n 10}"
+values 20 "{(n, m). tranclp succ n m}"
+
+inductive example_graph :: "int => int => bool"
+where
+ "example_graph 0 1"
+| "example_graph 1 2"
+| "example_graph 1 3"
+| "example_graph 4 7"
+| "example_graph 4 5"
+| "example_graph 5 6"
+| "example_graph 7 6"
+| "example_graph 7 8"
+
+inductive not_reachable_in_example_graph :: "int => int => bool"
+where "\<not> (tranclp example_graph x y) ==> not_reachable_in_example_graph x y"
+
+code_pred (expected_modes: i => i => bool) not_reachable_in_example_graph .
+
+thm not_reachable_in_example_graph.equation
+thm tranclp.equation
+value "not_reachable_in_example_graph 0 3"
+value "not_reachable_in_example_graph 4 8"
+value "not_reachable_in_example_graph 5 6"
+text {* rtrancl compilation is strange! *}
+(*
+value "not_reachable_in_example_graph 0 4"
+value "not_reachable_in_example_graph 1 6"
+value "not_reachable_in_example_graph 8 4"*)
+
+code_pred [dseq] not_reachable_in_example_graph .
+
+values [dseq 6] "{x. tranclp example_graph 0 3}"
+
+values [dseq 0] "{x. not_reachable_in_example_graph 0 3}"
+values [dseq 0] "{x. not_reachable_in_example_graph 0 4}"
+values [dseq 20] "{x. not_reachable_in_example_graph 0 4}"
+values [dseq 6] "{x. not_reachable_in_example_graph 0 3}"
+values [dseq 3] "{x. not_reachable_in_example_graph 4 2}"
+values [dseq 6] "{x. not_reachable_in_example_graph 4 2}"
+
+
+inductive not_reachable_in_example_graph' :: "int => int => bool"
+where "\<not> (rtranclp example_graph x y) ==> not_reachable_in_example_graph' x y"
+
+code_pred not_reachable_in_example_graph' .
+
+value "not_reachable_in_example_graph' 0 3"
+(* value "not_reachable_in_example_graph' 0 5" would not terminate *)
+
+
+(*values [depth_limited 0] "{x. not_reachable_in_example_graph' 0 3}"*)
+(*values [depth_limited 3] "{x. not_reachable_in_example_graph' 0 3}"*) (* fails with undefined *)
+(*values [depth_limited 5] "{x. not_reachable_in_example_graph' 0 3}"*)
+(*values [depth_limited 1] "{x. not_reachable_in_example_graph' 0 4}"*)
+(*values [depth_limit = 4] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
+(*values [depth_limit = 20] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
+
+code_pred [dseq] not_reachable_in_example_graph' .
+
+(*thm not_reachable_in_example_graph'.dseq_equation*)
+
+(*values [dseq 0] "{x. not_reachable_in_example_graph' 0 3}"*)
+(*values [depth_limited 3] "{x. not_reachable_in_example_graph' 0 3}"*) (* fails with undefined *)
+(*values [depth_limited 5] "{x. not_reachable_in_example_graph' 0 3}"
+values [depth_limited 1] "{x. not_reachable_in_example_graph' 0 4}"*)
+(*values [depth_limit = 4] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
+(*values [depth_limit = 20] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
+
+
+subsection {* IMP *}
+
+types
+ var = nat
+ state = "int list"
+
+datatype com =
+ Skip |
+ Ass var "state => int" |
+ Seq com com |
+ IF "state => bool" com com |
+ While "state => bool" com
+
+inductive exec :: "com => state => state => bool" where
+"exec Skip s s" |
+"exec (Ass x e) s (s[x := e(s)])" |
+"exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" |
+"b s ==> exec c1 s t ==> exec (IF b c1 c2) s t" |
+"~b s ==> exec c2 s t ==> exec (IF b c1 c2) s t" |
+"~b s ==> exec (While b c) s s" |
+"b s1 ==> exec c s1 s2 ==> exec (While b c) s2 s3 ==> exec (While b c) s1 s3"
+
+code_pred exec .
+
+values "{t. exec
+ (While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1))))
+ [3,5] t}"
+
+
+inductive tupled_exec :: "(com \<times> state \<times> state) \<Rightarrow> bool" where
+"tupled_exec (Skip, s, s)" |
+"tupled_exec (Ass x e, s, s[x := e(s)])" |
+"tupled_exec (c1, s1, s2) ==> tupled_exec (c2, s2, s3) ==> tupled_exec (Seq c1 c2, s1, s3)" |
+"b s ==> tupled_exec (c1, s, t) ==> tupled_exec (IF b c1 c2, s, t)" |
+"~b s ==> tupled_exec (c2, s, t) ==> tupled_exec (IF b c1 c2, s, t)" |
+"~b s ==> tupled_exec (While b c, s, s)" |
+"b s1 ==> tupled_exec (c, s1, s2) ==> tupled_exec (While b c, s2, s3) ==> tupled_exec (While b c, s1, s3)"
+
+code_pred tupled_exec .
+
+values "{s. tupled_exec (While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1))), [3, 5], s)}"
+
+subsection {* CCS *}
+
+text{* This example formalizes finite CCS processes without communication or
+recursion. For simplicity, labels are natural numbers. *}
+
+datatype proc = nil | pre nat proc | or proc proc | par proc proc
+
+inductive step :: "proc \<Rightarrow> nat \<Rightarrow> proc \<Rightarrow> bool" where
+"step (pre n p) n p" |
+"step p1 a q \<Longrightarrow> step (or p1 p2) a q" |
+"step p2 a q \<Longrightarrow> step (or p1 p2) a q" |
+"step p1 a q \<Longrightarrow> step (par p1 p2) a (par q p2)" |
+"step p2 a q \<Longrightarrow> step (par p1 p2) a (par p1 q)"
+
+code_pred step .
+
+inductive steps where
+"steps p [] p" |
+"step p a q \<Longrightarrow> steps q as r \<Longrightarrow> steps p (a#as) r"
+
+code_pred steps .
+
+values 3
+ "{as . steps (par (or (pre 0 nil) (pre 1 nil)) (pre 2 nil)) as (par nil nil)}"
+
+values 5
+ "{as . steps (par (or (pre 0 nil) (pre 1 nil)) (pre 2 nil)) as (par nil nil)}"
+
+values 3 "{(a,q). step (par nil nil) a q}"
+
+
+inductive tupled_step :: "(proc \<times> nat \<times> proc) \<Rightarrow> bool"
+where
+"tupled_step (pre n p, n, p)" |
+"tupled_step (p1, a, q) \<Longrightarrow> tupled_step (or p1 p2, a, q)" |
+"tupled_step (p2, a, q) \<Longrightarrow> tupled_step (or p1 p2, a, q)" |
+"tupled_step (p1, a, q) \<Longrightarrow> tupled_step (par p1 p2, a, par q p2)" |
+"tupled_step (p2, a, q) \<Longrightarrow> tupled_step (par p1 p2, a, par p1 q)"
+
+code_pred tupled_step .
+thm tupled_step.equation
+
+subsection {* divmod *}
+
+inductive divmod_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
+ "k < l \<Longrightarrow> divmod_rel k l 0 k"
+ | "k \<ge> l \<Longrightarrow> divmod_rel (k - l) l q r \<Longrightarrow> divmod_rel k l (Suc q) r"
+
+code_pred divmod_rel ..
+thm divmod_rel.equation
+value [code] "Predicate.the (divmod_rel_i_i_o_o 1705 42)"
+
+subsection {* Transforming predicate logic into logic programs *}
+
+subsection {* Transforming functions into logic programs *}
+definition
+ "case_f xs ys = (case (xs @ ys) of [] => [] | (x # xs) => xs)"
+
+code_pred [inductify] case_f .
+thm case_fP.equation
+thm case_fP.intros
+
+fun fold_map_idx where
+ "fold_map_idx f i y [] = (y, [])"
+| "fold_map_idx f i y (x # xs) =
+ (let (y', x') = f i y x; (y'', xs') = fold_map_idx f (Suc i) y' xs
+ in (y'', x' # xs'))"
+
+text {* mode analysis explores thousand modes - this is infeasible at the moment... *}
+(*code_pred [inductify, show_steps] fold_map_idx .*)
+
+subsection {* Minimum *}
+
+definition Min
+where "Min s r x \<equiv> s x \<and> (\<forall>y. r x y \<longrightarrow> x = y)"
+
+code_pred [inductify] Min .
+thm Min.equation
+
+subsection {* Lexicographic order *}
+
+declare lexord_def[code_pred_def]
+code_pred [inductify] lexord .
+code_pred [random_dseq inductify] lexord .
+
+thm lexord.equation
+thm lexord.random_dseq_equation
+
+inductive less_than_nat :: "nat * nat => bool"
+where
+ "less_than_nat (0, x)"
+| "less_than_nat (x, y) ==> less_than_nat (Suc x, Suc y)"
+
+code_pred less_than_nat .
+
+code_pred [dseq] less_than_nat .
+code_pred [random_dseq] less_than_nat .
+
+inductive test_lexord :: "nat list * nat list => bool"
+where
+ "lexord less_than_nat (xs, ys) ==> test_lexord (xs, ys)"
+
+code_pred test_lexord .
+code_pred [dseq] test_lexord .
+code_pred [random_dseq] test_lexord .
+thm test_lexord.dseq_equation
+thm test_lexord.random_dseq_equation
+
+values "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"
+(*values [depth_limited 5] "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"*)
+
+declare list.size(3,4)[code_pred_def]
+lemmas [code_pred_def] = lexn_conv lex_conv lenlex_conv
+(*
+code_pred [inductify] lexn .
+thm lexn.equation
+*)
+(*
+code_pred [random_dseq inductify] lexn .
+thm lexn.random_dseq_equation
+
+values [random_dseq 4, 4, 6] 100 "{(n, xs, ys::int list). lexn (%(x, y). x <= y) n (xs, ys)}"
+*)
+inductive has_length
+where
+ "has_length [] 0"
+| "has_length xs i ==> has_length (x # xs) (Suc i)"
+
+lemma has_length:
+ "has_length xs n = (length xs = n)"
+proof (rule iffI)
+ assume "has_length xs n"
+ from this show "length xs = n"
+ by (rule has_length.induct) auto
+next
+ assume "length xs = n"
+ from this show "has_length xs n"
+ by (induct xs arbitrary: n) (auto intro: has_length.intros)
+qed
+
+lemma lexn_intros [code_pred_intro]:
+ "has_length xs i ==> has_length ys i ==> r (x, y) ==> lexn r (Suc i) (x # xs, y # ys)"
+ "lexn r i (xs, ys) ==> lexn r (Suc i) (x # xs, x # ys)"
+proof -
+ assume "has_length xs i" "has_length ys i" "r (x, y)"
+ from this has_length show "lexn r (Suc i) (x # xs, y # ys)"
+ unfolding lexn_conv Collect_def mem_def
+ by fastsimp
+next
+ assume "lexn r i (xs, ys)"
+ thm lexn_conv
+ from this show "lexn r (Suc i) (x#xs, x#ys)"
+ unfolding Collect_def mem_def lexn_conv
+ apply auto
+ apply (rule_tac x="x # xys" in exI)
+ by auto
+qed
+
+code_pred [random_dseq inductify] lexn
+proof -
+ fix r n xs ys
+ assume 1: "lexn r n (xs, ys)"
+ assume 2: "\<And>r' i x xs' y ys'. r = r' ==> n = Suc i ==> (xs, ys) = (x # xs', y # ys') ==> has_length xs' i ==> has_length ys' i ==> r' (x, y) ==> thesis"
+ assume 3: "\<And>r' i x xs' ys'. r = r' ==> n = Suc i ==> (xs, ys) = (x # xs', x # ys') ==> lexn r' i (xs', ys') ==> thesis"
+ from 1 2 3 show thesis
+ unfolding lexn_conv Collect_def mem_def
+ apply (auto simp add: has_length)
+ apply (case_tac xys)
+ apply auto
+ apply fastsimp
+ apply fastsimp done
+qed
+
+
+values [random_dseq 1, 2, 5] 10 "{(n, xs, ys::int list). lexn (%(x, y). x <= y) n (xs, ys)}"
+thm lenlex_conv
+thm lex_conv
+declare list.size(3,4)[code_pred_def]
+(*code_pred [inductify, show_steps, show_intermediate_results] length .*)
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name Orderings.top_class.top}] *}
+code_pred [inductify] lex .
+thm lex.equation
+thm lex_def
+declare lenlex_conv[code_pred_def]
+code_pred [inductify] lenlex .
+thm lenlex.equation
+
+code_pred [random_dseq inductify] lenlex .
+thm lenlex.random_dseq_equation
+
+values [random_dseq 4, 2, 4] 100 "{(xs, ys::int list). lenlex (%(x, y). x <= y) (xs, ys)}"
+thm lists.intros
+
+code_pred [inductify] lists .
+thm lists.equation
+
+subsection {* AVL Tree *}
+
+datatype 'a tree = ET | MKT 'a "'a tree" "'a tree" nat
+fun height :: "'a tree => nat" where
+"height ET = 0"
+| "height (MKT x l r h) = max (height l) (height r) + 1"
+
+consts avl :: "'a tree => bool"
+primrec
+ "avl ET = True"
+ "avl (MKT x l r h) = ((height l = height r \<or> height l = 1 + height r \<or> height r = 1+height l) \<and>
+ h = max (height l) (height r) + 1 \<and> avl l \<and> avl r)"
+(*
+code_pred [inductify] avl .
+thm avl.equation*)
+
+code_pred [random_dseq inductify] avl .
+thm avl.random_dseq_equation
+
+values [random_dseq 2, 1, 7] 5 "{t:: int tree. avl t}"
+
+fun set_of
+where
+"set_of ET = {}"
+| "set_of (MKT n l r h) = insert n (set_of l \<union> set_of r)"
+
+fun is_ord :: "nat tree => bool"
+where
+"is_ord ET = True"
+| "is_ord (MKT n l r h) =
+ ((\<forall>n' \<in> set_of l. n' < n) \<and> (\<forall>n' \<in> set_of r. n < n') \<and> is_ord l \<and> is_ord r)"
+
+code_pred (expected_modes: i => o => bool, i => i => bool) [inductify] set_of .
+thm set_of.equation
+
+code_pred (expected_modes: i => bool) [inductify] is_ord .
+thm is_ord_aux.equation
+thm is_ord.equation
+
+
+subsection {* Definitions about Relations *}
+term "converse"
+code_pred (modes:
+ (i * i => bool) => i * i => bool,
+ (i * o => bool) => o * i => bool,
+ (i * o => bool) => i * i => bool,
+ (o * i => bool) => i * o => bool,
+ (o * i => bool) => i * i => bool,
+ (o * o => bool) => o * o => bool,
+ (o * o => bool) => i * o => bool,
+ (o * o => bool) => o * i => bool,
+ (o * o => bool) => i * i => bool) [inductify] converse .
+
+thm converse.equation
+code_pred [inductify] rel_comp .
+thm rel_comp.equation
+code_pred [inductify] Image .
+thm Image.equation
+declare singleton_iff[code_pred_inline]
+declare Id_on_def[unfolded Bex_def UNION_def singleton_iff, code_pred_def]
+
+code_pred (expected_modes:
+ (o => bool) => o => bool,
+ (o => bool) => i * o => bool,
+ (o => bool) => o * i => bool,
+ (o => bool) => i => bool,
+ (i => bool) => i * o => bool,
+ (i => bool) => o * i => bool,
+ (i => bool) => i => bool) [inductify] Id_on .
+thm Id_on.equation
+thm Domain_def
+code_pred (modes:
+ (o * o => bool) => o => bool,
+ (o * o => bool) => i => bool,
+ (i * o => bool) => i => bool) [inductify] Domain .
+thm Domain.equation
+
+thm Range_def
+code_pred (modes:
+ (o * o => bool) => o => bool,
+ (o * o => bool) => i => bool,
+ (o * i => bool) => i => bool) [inductify] Range .
+thm Range.equation
+
+code_pred [inductify] Field .
+thm Field.equation
+
+thm refl_on_def
+code_pred [inductify] refl_on .
+thm refl_on.equation
+code_pred [inductify] total_on .
+thm total_on.equation
+code_pred [inductify] antisym .
+thm antisym.equation
+code_pred [inductify] trans .
+thm trans.equation
+code_pred [inductify] single_valued .
+thm single_valued.equation
+thm inv_image_def
+code_pred [inductify] inv_image .
+thm inv_image.equation
+
+subsection {* Inverting list functions *}
+
+(*code_pred [inductify] length .
+code_pred [random inductify] length .
+thm size_listP.equation
+thm size_listP.random_equation
+*)
+(*values [random] 1 "{xs. size_listP (xs::nat list) (5::nat)}"*)
+
+code_pred (expected_modes: i => o => bool, o => i => bool, i => i => bool) [inductify] List.concat .
+thm concatP.equation
+
+values "{ys. concatP [[1, 2], [3, (4::int)]] ys}"
+values "{ys. concatP [[1, 2], [3]] [1, 2, (3::nat)]}"
+
+code_pred [dseq inductify] List.concat .
+thm concatP.dseq_equation
+
+values [dseq 3] 3
+ "{xs. concatP xs ([0] :: nat list)}"
+
+values [dseq 5] 3
+ "{xs. concatP xs ([1] :: int list)}"
+
+values [dseq 5] 3
+ "{xs. concatP xs ([1] :: nat list)}"
+
+values [dseq 5] 3
+ "{xs. concatP xs [(1::int), 2]}"
+
+code_pred (expected_modes: i => o => bool, i => i => bool) [inductify] hd .
+thm hdP.equation
+values "{x. hdP [1, 2, (3::int)] x}"
+values "{(xs, x). hdP [1, 2, (3::int)] 1}"
+
+code_pred (expected_modes: i => o => bool, i => i => bool) [inductify] tl .
+thm tlP.equation
+values "{x. tlP [1, 2, (3::nat)] x}"
+values "{x. tlP [1, 2, (3::int)] [3]}"
+
+code_pred [inductify] last .
+thm lastP.equation
+
+code_pred [inductify] butlast .
+thm butlastP.equation
+
+code_pred [inductify] take .
+thm takeP.equation
+
+code_pred [inductify] drop .
+thm dropP.equation
+code_pred [inductify] zip .
+thm zipP.equation
+
+code_pred [inductify] upt .
+code_pred [inductify] remdups .
+thm remdupsP.equation
+code_pred [dseq inductify] remdups .
+values [dseq 4] 5 "{xs. remdupsP xs [1, (2::int)]}"
+
+code_pred [inductify] remove1 .
+thm remove1P.equation
+values "{xs. remove1P 1 xs [2, (3::int)]}"
+
+code_pred [inductify] removeAll .
+thm removeAllP.equation
+code_pred [dseq inductify] removeAll .
+
+values [dseq 4] 10 "{xs. removeAllP 1 xs [(2::nat)]}"
+
+code_pred [inductify] distinct .
+thm distinct.equation
+code_pred [inductify] replicate .
+thm replicateP.equation
+values 5 "{(n, xs). replicateP n (0::int) xs}"
+
+code_pred [inductify] splice .
+thm splice.simps
+thm spliceP.equation
+
+values "{xs. spliceP xs [1, 2, 3] [1, 1, 1, 2, 1, (3::nat)]}"
+
+code_pred [inductify] List.rev .
+code_pred [inductify] map .
+code_pred [inductify] foldr .
+code_pred [inductify] foldl .
+code_pred [inductify] filter .
+code_pred [random_dseq inductify] filter .
+
+subsection {* Context Free Grammar *}
+
+datatype alphabet = a | b
+
+inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
+ "[] \<in> S\<^isub>1"
+| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
+| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
+| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
+| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
+| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
+
+code_pred [inductify] S\<^isub>1p .
+code_pred [random_dseq inductify] S\<^isub>1p .
+thm S\<^isub>1p.equation
+thm S\<^isub>1p.random_dseq_equation
+
+values [random_dseq 5, 5, 5] 5 "{x. S\<^isub>1p x}"
+
+inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
+ "[] \<in> S\<^isub>2"
+| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
+| "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2"
+| "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
+| "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
+| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
+
+code_pred [random_dseq inductify] S\<^isub>2p .
+thm S\<^isub>2p.random_dseq_equation
+thm A\<^isub>2p.random_dseq_equation
+thm B\<^isub>2p.random_dseq_equation
+
+values [random_dseq 5, 5, 5] 10 "{x. S\<^isub>2p x}"
+
+inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
+ "[] \<in> S\<^isub>3"
+| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3"
+| "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3"
+| "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3"
+| "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3"
+| "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3"
+
+code_pred [inductify] S\<^isub>3p .
+thm S\<^isub>3p.equation
+
+values 10 "{x. S\<^isub>3p x}"
+
+inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
+ "[] \<in> S\<^isub>4"
+| "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4"
+| "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4"
+| "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4"
+| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
+| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
+| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
+
+code_pred (expected_modes: o => bool, i => bool) S\<^isub>4p .
+
+subsection {* Lambda *}
+
+datatype type =
+ Atom nat
+ | Fun type type (infixr "\<Rightarrow>" 200)
+
+datatype dB =
+ Var nat
+ | App dB dB (infixl "\<degree>" 200)
+ | Abs type dB
+
+primrec
+ nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
+where
+ "[]\<langle>i\<rangle> = None"
+| "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
+
+inductive nth_el' :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool"
+where
+ "nth_el' (x # xs) 0 x"
+| "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y"
+
+inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" ("_ \<turnstile> _ : _" [50, 50, 50] 50)
+ where
+ Var [intro!]: "nth_el' env x T \<Longrightarrow> env \<turnstile> Var x : T"
+ | Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)"
+ | App [intro!]: "env \<turnstile> s : T \<Rightarrow> U \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U"
+
+primrec
+ lift :: "[dB, nat] => dB"
+where
+ "lift (Var i) k = (if i < k then Var i else Var (i + 1))"
+ | "lift (s \<degree> t) k = lift s k \<degree> lift t k"
+ | "lift (Abs T s) k = Abs T (lift s (k + 1))"
+
+primrec
+ subst :: "[dB, dB, nat] => dB" ("_[_'/_]" [300, 0, 0] 300)
+where
+ subst_Var: "(Var i)[s/k] =
+ (if k < i then Var (i - 1) else if i = k then s else Var i)"
+ | subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
+ | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"
+
+inductive beta :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<beta>" 50)
+ where
+ beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
+ | appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
+ | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
+ | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"
+
+code_pred (expected_modes: i => i => o => bool, i => i => i => bool) typing .
+thm typing.equation
+
+code_pred (modes: i => i => bool, i => o => bool as reduce') beta .
+thm beta.equation
+
+values "{x. App (Abs (Atom 0) (Var 0)) (Var 1) \<rightarrow>\<^sub>\<beta> x}"
+
+definition "reduce t = Predicate.the (reduce' t)"
+
+value "reduce (App (Abs (Atom 0) (Var 0)) (Var 1))"
+
+code_pred [dseq] typing .
+code_pred [random_dseq] typing .
+
+values [random_dseq 1,1,5] 10 "{(\<Gamma>, t, T). \<Gamma> \<turnstile> t : T}"
+
+subsection {* A minimal example of yet another semantics *}
+
+text {* thanks to Elke Salecker *}
+
+types
+ vname = nat
+ vvalue = int
+ var_assign = "vname \<Rightarrow> vvalue" --"variable assignment"
+
+datatype ir_expr =
+ IrConst vvalue
+| ObjAddr vname
+| Add ir_expr ir_expr
+
+datatype val =
+ IntVal vvalue
+
+record configuration =
+ Env :: var_assign
+
+inductive eval_var ::
+ "ir_expr \<Rightarrow> configuration \<Rightarrow> val \<Rightarrow> bool"
+where
+ irconst: "eval_var (IrConst i) conf (IntVal i)"
+| objaddr: "\<lbrakk> Env conf n = i \<rbrakk> \<Longrightarrow> eval_var (ObjAddr n) conf (IntVal i)"
+| plus: "\<lbrakk> eval_var l conf (IntVal vl); eval_var r conf (IntVal vr) \<rbrakk> \<Longrightarrow> eval_var (Add l r) conf (IntVal (vl+vr))"
+
+
+code_pred eval_var .
+thm eval_var.equation
+
+values "{val. eval_var (Add (IrConst 1) (IrConst 2)) (| Env = (\<lambda>x. 0)|) val}"
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy Thu Mar 25 17:56:31 2010 +0100
@@ -0,0 +1,337 @@
+theory Predicate_Compile_Quickcheck_Examples
+imports Predicate_Compile_Quickcheck
+begin
+
+section {* Sets *}
+
+lemma "x \<in> {(1::nat)} ==> False"
+quickcheck[generator=predicate_compile_wo_ff, iterations=10]
+oops
+
+lemma "x \<in> {Suc 0, Suc (Suc 0)} ==> x \<noteq> Suc 0"
+quickcheck[generator=predicate_compile_wo_ff]
+oops
+
+lemma "x \<in> {Suc 0, Suc (Suc 0)} ==> x = Suc 0"
+quickcheck[generator=predicate_compile_wo_ff]
+oops
+
+lemma "x \<in> {Suc 0, Suc (Suc 0)} ==> x <= Suc 0"
+quickcheck[generator=predicate_compile_wo_ff]
+oops
+
+section {* Numerals *}
+
+lemma
+ "x \<in> {1, 2, (3::nat)} ==> x = 1 \<or> x = 2"
+quickcheck[generator=predicate_compile_wo_ff]
+oops
+
+lemma "x \<in> {1, 2, (3::nat)} ==> x < 3"
+quickcheck[generator=predicate_compile_wo_ff]
+oops
+
+lemma
+ "x \<in> {1, 2} \<union> {3, 4} ==> x = (1::nat) \<or> x = (2::nat)"
+quickcheck[generator=predicate_compile_wo_ff]
+oops
+
+section {* Context Free Grammar *}
+
+datatype alphabet = a | b
+
+inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
+ "[] \<in> S\<^isub>1"
+| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
+| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
+| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
+| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
+| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
+
+lemma
+ "w \<in> S\<^isub>1 \<Longrightarrow> w = []"
+quickcheck[generator = predicate_compile_ff_nofs, iterations=1]
+oops
+
+theorem S\<^isub>1_sound:
+"w \<in> S\<^isub>1 \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+quickcheck[generator=predicate_compile_ff_nofs, size=15]
+oops
+
+
+inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
+ "[] \<in> S\<^isub>2"
+| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
+| "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2"
+| "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
+| "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
+| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
+(*
+code_pred [random_dseq inductify] S\<^isub>2 .
+thm S\<^isub>2.random_dseq_equation
+thm A\<^isub>2.random_dseq_equation
+thm B\<^isub>2.random_dseq_equation
+
+values [random_dseq 1, 2, 8] 10 "{x. S\<^isub>2 x}"
+
+lemma "w \<in> S\<^isub>2 ==> w \<noteq> [] ==> w \<noteq> [b, a] ==> w \<in> {}"
+quickcheck[generator=predicate_compile, size=8]
+oops
+
+lemma "[x <- w. x = a] = []"
+quickcheck[generator=predicate_compile]
+oops
+
+declare list.size(3,4)[code_pred_def]
+
+(*
+lemma "length ([x \<leftarrow> w. x = a]) = (0::nat)"
+quickcheck[generator=predicate_compile]
+oops
+*)
+
+lemma
+"w \<in> S\<^isub>2 ==> length [x \<leftarrow> w. x = a] <= Suc (Suc 0)"
+quickcheck[generator=predicate_compile, size = 10, iterations = 1]
+oops
+*)
+theorem S\<^isub>2_sound:
+"w \<in> S\<^isub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+quickcheck[generator=predicate_compile_ff_nofs, size=5, iterations=10]
+oops
+
+inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
+ "[] \<in> S\<^isub>3"
+| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3"
+| "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3"
+| "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3"
+| "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3"
+| "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3"
+
+code_pred [inductify] S\<^isub>3 .
+thm S\<^isub>3.equation
+(*
+values 10 "{x. S\<^isub>3 x}"
+*)
+
+
+lemma S\<^isub>3_sound:
+"w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+quickcheck[generator=predicate_compile_ff_fs, size=10, iterations=10]
+oops
+
+lemma "\<not> (length w > 2) \<or> \<not> (length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b])"
+quickcheck[size=10, generator = predicate_compile_ff_fs]
+oops
+
+theorem S\<^isub>3_complete:
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. b = x] \<longrightarrow> w \<in> S\<^isub>3"
+(*quickcheck[generator=SML]*)
+quickcheck[generator=predicate_compile_ff_fs, size=10, iterations=100]
+oops
+
+
+inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
+ "[] \<in> S\<^isub>4"
+| "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4"
+| "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4"
+| "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4"
+| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
+| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
+| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
+
+theorem S\<^isub>4_sound:
+"w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+quickcheck[generator = predicate_compile_ff_nofs, size=5, iterations=1]
+oops
+
+theorem S\<^isub>4_complete:
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4"
+quickcheck[generator = predicate_compile_ff_nofs, size=5, iterations=1]
+oops
+
+hide const a b
+
+subsection {* Lexicographic order *}
+(* TODO *)
+(*
+lemma
+ "(u, v) : lexord r ==> (x @ u, y @ v) : lexord r"
+oops
+*)
+subsection {* IMP *}
+
+types
+ var = nat
+ state = "int list"
+
+datatype com =
+ Skip |
+ Ass var "int" |
+ Seq com com |
+ IF "state list" com com |
+ While "state list" com
+
+inductive exec :: "com => state => state => bool" where
+ "exec Skip s s" |
+ "exec (Ass x e) s (s[x := e])" |
+ "exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" |
+ "s \<in> set b ==> exec c1 s t ==> exec (IF b c1 c2) s t" |
+ "s \<notin> set b ==> exec c2 s t ==> exec (IF b c1 c2) s t" |
+ "s \<notin> set b ==> exec (While b c) s s" |
+ "s1 \<in> set b ==> exec c s1 s2 ==> exec (While b c) s2 s3 ==> exec (While b c) s1 s3"
+
+code_pred [random_dseq] exec .
+
+values [random_dseq 1, 2, 3] 10 "{(c, s, s'). exec c s s'}"
+
+lemma
+ "exec c s s' ==> exec (Seq c c) s s'"
+(*quickcheck[generator = predicate_compile_wo_ff, size=2, iterations=10]*)
+oops
+
+subsection {* Lambda *}
+
+datatype type =
+ Atom nat
+ | Fun type type (infixr "\<Rightarrow>" 200)
+
+datatype dB =
+ Var nat
+ | App dB dB (infixl "\<degree>" 200)
+ | Abs type dB
+
+primrec
+ nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
+where
+ "[]\<langle>i\<rangle> = None"
+| "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
+
+inductive nth_el' :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool"
+where
+ "nth_el' (x # xs) 0 x"
+| "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y"
+
+inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" ("_ \<turnstile> _ : _" [50, 50, 50] 50)
+ where
+ Var [intro!]: "nth_el' env x T \<Longrightarrow> env \<turnstile> Var x : T"
+ | Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)"
+ | App [intro!]: "env \<turnstile> s : U \<Rightarrow> T \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U"
+
+primrec
+ lift :: "[dB, nat] => dB"
+where
+ "lift (Var i) k = (if i < k then Var i else Var (i + 1))"
+ | "lift (s \<degree> t) k = lift s k \<degree> lift t k"
+ | "lift (Abs T s) k = Abs T (lift s (k + 1))"
+
+primrec
+ subst :: "[dB, dB, nat] => dB" ("_[_'/_]" [300, 0, 0] 300)
+where
+ subst_Var: "(Var i)[s/k] =
+ (if k < i then Var (i - 1) else if i = k then s else Var i)"
+ | subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
+ | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"
+
+inductive beta :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<beta>" 50)
+ where
+ beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
+ | appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
+ | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
+ | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"
+
+lemma
+ "\<Gamma> \<turnstile> t : U \<Longrightarrow> t \<rightarrow>\<^sub>\<beta> t' \<Longrightarrow> \<Gamma> \<turnstile> t' : U"
+quickcheck[generator = predicate_compile_ff_fs, size = 7, iterations = 10]
+oops
+
+subsection {* JAD *}
+
+definition matrix :: "('a :: semiring_0) list list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
+ "matrix M rs cs \<longleftrightarrow> (\<forall> row \<in> set M. length row = cs) \<and> length M = rs"
+(*
+code_pred [random_dseq inductify] matrix .
+thm matrix.random_dseq_equation
+
+thm matrix_aux.random_dseq_equation
+
+values [random_dseq 3, 2] 10 "{(M, rs, cs). matrix (M:: int list list) rs cs}"
+*)
+lemma [code_pred_intro]:
+ "matrix [] 0 m"
+ "matrix xss n m ==> length xs = m ==> matrix (xs # xss) (Suc n) m"
+proof -
+ show "matrix [] 0 m" unfolding matrix_def by auto
+next
+ show "matrix xss n m ==> length xs = m ==> matrix (xs # xss) (Suc n) m"
+ unfolding matrix_def by auto
+qed
+
+code_pred [random_dseq inductify] matrix
+ apply (cases x)
+ unfolding matrix_def apply fastsimp
+ apply fastsimp done
+
+
+values [random_dseq 2, 2, 15] 6 "{(M::int list list, n, m). matrix M n m}"
+
+definition "scalar_product v w = (\<Sum> (x, y)\<leftarrow>zip v w. x * y)"
+
+definition mv :: "('a \<Colon> semiring_0) list list \<Rightarrow> 'a list \<Rightarrow> 'a list"
+ where [simp]: "mv M v = map (scalar_product v) M"
+text {*
+ This defines the matrix vector multiplication. To work properly @{term
+"matrix M m n \<and> length v = n"} must hold.
+*}
+
+subsection "Compressed matrix"
+
+definition "sparsify xs = [i \<leftarrow> zip [0..<length xs] xs. snd i \<noteq> 0]"
+(*
+lemma sparsify_length: "(i, x) \<in> set (sparsify xs) \<Longrightarrow> i < length xs"
+ by (auto simp: sparsify_def set_zip)
+
+lemma listsum_sparsify[simp]:
+ fixes v :: "('a \<Colon> semiring_0) list"
+ assumes "length w = length v"
+ shows "(\<Sum>x\<leftarrow>sparsify w. (\<lambda>(i, x). v ! i) x * snd x) = scalar_product v w"
+ (is "(\<Sum>x\<leftarrow>_. ?f x) = _")
+ unfolding sparsify_def scalar_product_def
+ using assms listsum_map_filter[where f="?f" and P="\<lambda> i. snd i \<noteq> (0::'a)"]
+ by (simp add: listsum_setsum)
+*)
+definition [simp]: "unzip w = (map fst w, map snd w)"
+
+primrec insert :: "('a \<Rightarrow> 'b \<Colon> linorder) => 'a \<Rightarrow> 'a list => 'a list" where
+ "insert f x [] = [x]" |
+ "insert f x (y # ys) = (if f y < f x then y # insert f x ys else x # y # ys)"
+
+primrec sort :: "('a \<Rightarrow> 'b \<Colon> linorder) \<Rightarrow> 'a list => 'a list" where
+ "sort f [] = []" |
+ "sort f (x # xs) = insert f x (sort f xs)"
+
+definition
+ "length_permutate M = (unzip o sort (length o snd)) (zip [0 ..< length M] M)"
+(*
+definition
+ "transpose M = [map (\<lambda> xs. xs ! i) (takeWhile (\<lambda> xs. i < length xs) M). i \<leftarrow> [0 ..< length (M ! 0)]]"
+*)
+definition
+ "inflate upds = foldr (\<lambda> (i, x) upds. upds[i := x]) upds (replicate (length upds) 0)"
+
+definition
+ "jad = apsnd transpose o length_permutate o map sparsify"
+
+definition
+ "jad_mv v = inflate o split zip o apsnd (map listsum o transpose o map (map (\<lambda> (i, x). v ! i * x)))"
+
+lemma "matrix (M::int list list) rs cs \<Longrightarrow> False"
+quickcheck[generator = predicate_compile_ff_nofs, size = 6]
+oops
+
+lemma
+ "\<lbrakk> matrix M rs cs ; length v = cs \<rbrakk> \<Longrightarrow> jad_mv v (jad M) = mv M v"
+quickcheck[generator = predicate_compile_wo_ff]
+oops
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Predicate_Compile_Examples/ROOT.ML Thu Mar 25 17:56:31 2010 +0100
@@ -0,0 +1,1 @@
+use_thys ["Predicate_Compile_Examples", "Predicate_Compile_Quickcheck_Examples"];
--- a/src/HOL/Probability/Probability_Space.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOL/Probability/Probability_Space.thy Thu Mar 25 17:56:31 2010 +0100
@@ -38,7 +38,7 @@
definition
"conditional_prob e1 e2 \<equiv> conditional_expectation (indicator_fn e1) e2"
-lemma positive: "positive M prob"
+lemma positive': "positive M prob"
unfolding positive_def using positive empty_measure by blast
lemma prob_compl:
@@ -60,7 +60,7 @@
lemma prob_space_increasing:
"increasing M prob"
-by (rule additive_increasing[OF positive additive])
+by (rule additive_increasing[OF positive' additive])
lemma prob_subadditive:
assumes "s \<in> events" "t \<in> events"
@@ -167,7 +167,7 @@
assms rf' unfolding f'_def by blast
hence absinc: "\<And> i. \<bar> prob (f' i) \<bar> \<le> prob (f i)"
- using abs_of_nonneg positive[unfolded positive_def]
+ using abs_of_nonneg positive'[unfolded positive_def]
assms rf' by auto
have "prob (\<Union> i. f i) = prob (\<Union> i. f' i)" using uf'f by simp
@@ -192,7 +192,7 @@
using assms
proof -
have leq0: "0 \<le> prob (\<Union> i. c i)"
- using assms positive[unfolded positive_def, rule_format]
+ using assms positive'[unfolded positive_def, rule_format]
by auto
have "prob (\<Union> i. c i) \<le> (\<Sum> i. prob (c i))"
@@ -285,7 +285,7 @@
have pos: "\<And> e. e \<in> sets s \<Longrightarrow> distribution X e \<ge> 0"
unfolding distribution_def
- using positive[unfolded positive_def]
+ using positive'[unfolded positive_def]
assms[unfolded measurable_def] by auto
have cas: "countably_additive ?N (distribution X)"
@@ -366,7 +366,7 @@
let "?M" = "\<lparr> space = ((X ` (space M)) \<times> (Y ` (space M))),
sets = Pow ((X ` (space M)) \<times> (Y ` (space M)))\<rparr>"
have pos: "positive ?S (\<lambda>a. prob ((\<lambda>x. (X x,Y x)) -` a \<inter> space M))"
- unfolding positive_def using positive[unfolded positive_def] assms by auto
+ unfolding positive_def using positive'[unfolded positive_def] assms by auto
{ fix x y
have A: "((\<lambda>x. (X x, Y x)) -` x) \<inter> space M \<in> sets M" using assms by auto
have B: "((\<lambda>x. (X x, Y x)) -` y) \<inter> space M \<in> sets M" using assms by auto
@@ -397,7 +397,7 @@
proof -
let "?S" = "\<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2) \<rparr>"
let "?M" = "\<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2), measure = joint_distribution X Y \<rparr>"
- have pos: "positive ?S (joint_distribution X Y)" using positive
+ have pos: "positive ?S (joint_distribution X Y)" using positive'
unfolding positive_def joint_distribution_def using assms by auto
{ fix x y
have A: "((\<lambda>x. (X x, Y x)) -` x) \<inter> space M \<in> sets M" using assms by auto
--- a/src/HOL/SMT/Examples/SMT_Examples.certs Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOL/SMT/Examples/SMT_Examples.certs Thu Mar 25 17:56:31 2010 +0100
@@ -1,4 +1,4 @@
-Fg1W6egjwo9zhhAmUXOW+w 8 0
+bb06851c317eb8b672e27364b0ae34a4e39eb880 8 0
#2 := false
#1 := true
#4 := (not true)
@@ -7,7 +7,7 @@
#20 := [asserted]: #4
[mp #20 #22]: false
unsat
-2+cndY9nzS72l7VvBCGRAw 19 0
+70d1f77bec207467bc0306af0d98a71fa8328274 19 0
#2 := false
decl up_1 :: bool
#4 := up_1
@@ -27,7 +27,7 @@
#23 := [asserted]: #7
[mp #23 #32]: false
unsat
-0vJQrobUDcQ9PkGJO8aM8g 25 0
+148012a9e9d44fe30a0c79e3344bdb805124f661 25 0
#2 := false
decl up_1 :: bool
#4 := up_1
@@ -53,7 +53,7 @@
#23 := [asserted]: #7
[mp #23 #38]: false
unsat
-AGGnpwEv208Vqxly7wTWHA 38 0
+bc9a25b7f6dc3ac2431ee71b6e71c5a7b25e89d1 38 0
#2 := false
decl up_2 :: bool
#5 := up_2
@@ -92,9 +92,9 @@
#30 := [and-elim #31]: #6
[mp #30 #52]: false
unsat
-wakXeIy1uoPgglzOQGFhJQ 1 0
-unsat
-cpSlDe0l7plVktRNxGU5dA 71 0
+9b3db6ce34c8a1806160f1349b898b6c5ca40ba0 1 0
+unsat
+912e9b7fb52f4a71d232354b3bb53c11e5a41ccd 71 0
#2 := false
decl up_1 :: bool
#4 := up_1
@@ -166,7 +166,7 @@
#31 := [asserted]: #15
[mp #31 #82]: false
unsat
-pg19mjJfV75T2QDrgWd4JA 57 0
+4d063d3cdf6657ddb4258379f900ef18e9042978 57 0
#2 := false
decl up_1 :: bool
#4 := up_1
@@ -224,7 +224,7 @@
#30 := [asserted]: #14
[mp #30 #70]: false
unsat
-Mj1B8X1MaN7xU/W4Kz3FoQ 194 0
+c6eb111aa830c9669a030c2e58b4e827706981da 194 0
#2 := false
decl up_1 :: bool
#4 := up_1
@@ -419,7 +419,7 @@
#237 := [mp #83 #236]: #75
[mp #237 #247]: false
unsat
-JkhYJB8FDavTZkizO1/9IA 52 0
+42890f9fa7c18237798ca55d0cf9dfff6f2f868a 52 0
#2 := false
decl uf_1 :: (-> T1 T1 T1)
decl uf_2 :: T1
@@ -472,7 +472,7 @@
#113 := [quant-inst]: #199
[unit-resolution #113 #536 #49]: false
unsat
-0ZdSZH2DbtjHNTyrDkZmXg 1667 0
+94169ba151f7a720e818287ce490015dde764bad 1667 0
#2 := false
decl up_54 :: bool
#126 := up_54
@@ -2140,7 +2140,7 @@
#2371 := [unit-resolution #891 #2369]: #166
[unit-resolution #2159 #2371 #2370 #2358 #2357]: false
unsat
-R3pmBDBlU9XdUrxJXhj7nA 78 0
+d8841d120b7cf772be783d793f759fb6353b9fcd 78 0
#2 := false
decl up_1 :: (-> int bool)
decl ?x1!0 :: int
@@ -2219,7 +2219,7 @@
#82 := [and-elim #81]: #55
[unit-resolution #82 #95]: false
unsat
-IBRj/loev6P6r0J+HOit6A 135 0
+18bde40beb0dd1fc2613a67805e24d767dd9a3bf 135 0
#2 := false
decl up_1 :: (-> T1 T2 bool)
#5 := (:var 0 T2)
@@ -2355,7 +2355,7 @@
#176 := [quant-inst]: #538
[unit-resolution #176 #252 #535]: false
unsat
-72504KVBixGB/87pOYiU/A 135 2
+013861f683c286a3ef38681266913846a3a39b9a 135 2
#2 := false
decl up_1 :: (-> T1 T2 bool)
#5 := (:var 0 T2)
@@ -2493,7 +2493,7 @@
unsat
WARNING: failed to find a pattern for quantifier (quantifier id: k!12)
-RaQLz4GxtUICnOD5WoYnzQ 56 0
+0e958e27514643bb596851e6dbb61a23f6b348b0 56 0
#2 := false
decl up_1 :: (-> T1 bool)
decl uf_2 :: T1
@@ -2550,7 +2550,7 @@
#120 := [quant-inst]: #206
[unit-resolution #120 #542 #41]: false
unsat
-NPQIgVPhSpgSLeS+u/EatQ 17 0
+6ecefa4023d224e6c51226d5bee17e2a19cc4333 17 0
#2 := false
#4 := 3::int
#5 := (= 3::int 3::int)
@@ -2568,7 +2568,7 @@
#22 := [asserted]: #6
[mp #22 #31]: false
unsat
-Lc9NwVtwY2Wo0G7UbFD1oA 17 0
+5e0256133fc82f0e2fea6597b863483e4e61d3c6 17 0
#2 := false
#4 := 3::real
#5 := (= 3::real 3::real)
@@ -2586,7 +2586,7 @@
#22 := [asserted]: #6
[mp #22 #31]: false
unsat
-pYVrUflpYrrZEWALJDnvPw 26 0
+55cf32b061b843ac5bcaefb74005a7dd3a24386f 26 0
#2 := false
#7 := 4::int
#5 := 1::int
@@ -2613,7 +2613,7 @@
#25 := [asserted]: #9
[mp #25 #40]: false
unsat
-FIqzVlbN8RT0iWarmBEpjw 41 0
+e81d17ec85af9db5ec6ba5bf4ced62daaa719ef3 41 0
#2 := false
decl uf_1 :: int
#4 := uf_1
@@ -2655,7 +2655,7 @@
#28 := [asserted]: #12
[mp #28 #52]: false
unsat
-HWVNtxMa8xktQsg8pHG+1w 35 0
+448f188ebf9d7fbd2920c0a51a8f105192e6af1a 35 0
#2 := false
#5 := 3::int
#6 := 8::int
@@ -2691,7 +2691,7 @@
#26 := [asserted]: #10
[mp #26 #51]: false
unsat
-M71YYpEc8u/aEIH3MOQrcg 250 0
+c3751ecae7701923f4ba6a90c6c6eee35ee1b13d 250 0
#2 := false
#7 := 0::real
decl uf_2 :: real
@@ -2942,7 +2942,7 @@
#294 := [unit-resolution #190 #293]: #188
[th-lemma #280 #294]: false
unsat
-G00bTqBjtW66EmwIZbXbOg 124 0
+8e6af261ea9f94b967813d4e2f2abcb94463d612 124 0
#2 := false
decl uf_1 :: (-> T1 T2)
decl uf_3 :: T1
@@ -3067,7 +3067,7 @@
#34 := [asserted]: #11
[unit-resolution #34 #536]: false
unsat
-6QdzkSy/RtEjUu+wUKIKqA 54 0
+243524c591f6dcfe16a79ddd249c64a337ff3612 54 0
#2 := false
#9 := 1::int
decl uf_1 :: int
@@ -3122,7 +3122,7 @@
#28 := [asserted]: #12
[mp #28 #67]: false
unsat
-xoSwaSeELbR0PHe0zb/BSg 63 0
+adfe7d6c2da6653191952bd9673c1274f94c2ab2 63 0
#2 := false
#11 := 0::int
decl uf_2 :: int
@@ -3186,7 +3186,7 @@
#76 := [mp #52 #75]: #63
[mp #76 #84]: false
unsat
-ciHqmDSmPpA15rO932dhvA 35 0
+3440e29713ba625633b10a2c4fdc186cb6e0cf3e 35 0
#2 := false
#6 := 5::int
#4 := 2::int
@@ -3222,7 +3222,7 @@
#25 := [asserted]: #9
[mp #25 #49]: false
unsat
-HzwFy7SRHqpspkYnzyeF4w 45 0
+5f3503ae4a43341932052006638f286bea551e87 45 0
#2 := false
#11 := 4::real
decl uf_2 :: real
@@ -3268,7 +3268,7 @@
#60 := [mp #36 #59]: #51
[th-lemma #60 #47 #38]: false
unsat
-XW7QIWmzYjfQXaHHPc98eA 59 0
+347776ce17d7d3f6d1252ead05795e4eee6f4b20 59 0
#2 := false
#16 := (not false)
decl uf_2 :: int
@@ -3328,94 +3328,7 @@
#34 := [asserted]: #18
[mp #34 #71]: false
unsat
-ZGL00TLLioiLlWFiXUnbxg 86 0
-#2 := false
-decl uf_1 :: int
-#5 := uf_1
-#7 := 2::int
-#29 := (* 2::int uf_1)
-#4 := 0::int
-#54 := (= 0::int #29)
-#55 := (not #54)
-#61 := (= #29 0::int)
-#104 := (not #61)
-#110 := (iff #104 #55)
-#108 := (iff #61 #54)
-#109 := [commutativity]: #108
-#111 := [monotonicity #109]: #110
-#62 := (<= #29 0::int)
-#100 := (not #62)
-#30 := (<= uf_1 0::int)
-#31 := (not #30)
-#6 := (< 0::int uf_1)
-#32 := (iff #6 #31)
-#33 := [rewrite]: #32
-#27 := [asserted]: #6
-#34 := [mp #27 #33]: #31
-#101 := (or #100 #30)
-#102 := [th-lemma]: #101
-#103 := [unit-resolution #102 #34]: #100
-#105 := (or #104 #62)
-#106 := [th-lemma]: #105
-#107 := [unit-resolution #106 #103]: #104
-#112 := [mp #107 #111]: #55
-#56 := (= uf_1 #29)
-#57 := (not #56)
-#53 := (= 0::int uf_1)
-#50 := (not #53)
-#58 := (and #50 #55 #57)
-#69 := (not #58)
-#42 := (distinct 0::int uf_1 #29)
-#47 := (not #42)
-#9 := (- uf_1 uf_1)
-#8 := (* uf_1 2::int)
-#10 := (distinct uf_1 #8 #9)
-#11 := (not #10)
-#48 := (iff #11 #47)
-#45 := (iff #10 #42)
-#39 := (distinct uf_1 #29 0::int)
-#43 := (iff #39 #42)
-#44 := [rewrite]: #43
-#40 := (iff #10 #39)
-#37 := (= #9 0::int)
-#38 := [rewrite]: #37
-#35 := (= #8 #29)
-#36 := [rewrite]: #35
-#41 := [monotonicity #36 #38]: #40
-#46 := [trans #41 #44]: #45
-#49 := [monotonicity #46]: #48
-#28 := [asserted]: #11
-#52 := [mp #28 #49]: #47
-#80 := (or #42 #69)
-#81 := [def-axiom]: #80
-#82 := [unit-resolution #81 #52]: #69
-#59 := (= uf_1 0::int)
-#83 := (not #59)
-#89 := (iff #83 #50)
-#87 := (iff #59 #53)
-#88 := [commutativity]: #87
-#90 := [monotonicity #88]: #89
-#84 := (or #83 #30)
-#85 := [th-lemma]: #84
-#86 := [unit-resolution #85 #34]: #83
-#91 := [mp #86 #90]: #50
-#64 := -1::int
-#65 := (* -1::int #29)
-#66 := (+ uf_1 #65)
-#68 := (>= #66 0::int)
-#92 := (not #68)
-#93 := (or #92 #30)
-#94 := [th-lemma]: #93
-#95 := [unit-resolution #94 #34]: #92
-#96 := (or #57 #68)
-#97 := [th-lemma]: #96
-#98 := [unit-resolution #97 #95]: #57
-#76 := (or #58 #53 #54 #56)
-#77 := [def-axiom]: #76
-#99 := [unit-resolution #77 #98 #91 #82]: #54
-[unit-resolution #99 #112]: false
-unsat
-DWt5rIK6NWlI4vrw+691Zg 212 0
+f7b1b99e3470f713e48aaae1336ace10223be6f0 212 0
#2 := false
decl uf_4 :: T1
#13 := uf_4
@@ -3628,7 +3541,94 @@
#519 := [unit-resolution #521 #518]: #547
[unit-resolution #519 #522]: false
unsat
-PaSeDRf7Set5ywlblDOoTg 673 0
+bf36938883aa38907d4d00c1860a1d18e7b620d0 86 0
+#2 := false
+decl uf_1 :: int
+#5 := uf_1
+#7 := 2::int
+#29 := (* 2::int uf_1)
+#4 := 0::int
+#54 := (= 0::int #29)
+#55 := (not #54)
+#61 := (= #29 0::int)
+#104 := (not #61)
+#110 := (iff #104 #55)
+#108 := (iff #61 #54)
+#109 := [commutativity]: #108
+#111 := [monotonicity #109]: #110
+#62 := (<= #29 0::int)
+#100 := (not #62)
+#30 := (<= uf_1 0::int)
+#31 := (not #30)
+#6 := (< 0::int uf_1)
+#32 := (iff #6 #31)
+#33 := [rewrite]: #32
+#27 := [asserted]: #6
+#34 := [mp #27 #33]: #31
+#101 := (or #100 #30)
+#102 := [th-lemma]: #101
+#103 := [unit-resolution #102 #34]: #100
+#105 := (or #104 #62)
+#106 := [th-lemma]: #105
+#107 := [unit-resolution #106 #103]: #104
+#112 := [mp #107 #111]: #55
+#56 := (= uf_1 #29)
+#57 := (not #56)
+#53 := (= 0::int uf_1)
+#50 := (not #53)
+#58 := (and #50 #55 #57)
+#69 := (not #58)
+#42 := (distinct 0::int uf_1 #29)
+#47 := (not #42)
+#9 := (- uf_1 uf_1)
+#8 := (* uf_1 2::int)
+#10 := (distinct uf_1 #8 #9)
+#11 := (not #10)
+#48 := (iff #11 #47)
+#45 := (iff #10 #42)
+#39 := (distinct uf_1 #29 0::int)
+#43 := (iff #39 #42)
+#44 := [rewrite]: #43
+#40 := (iff #10 #39)
+#37 := (= #9 0::int)
+#38 := [rewrite]: #37
+#35 := (= #8 #29)
+#36 := [rewrite]: #35
+#41 := [monotonicity #36 #38]: #40
+#46 := [trans #41 #44]: #45
+#49 := [monotonicity #46]: #48
+#28 := [asserted]: #11
+#52 := [mp #28 #49]: #47
+#80 := (or #42 #69)
+#81 := [def-axiom]: #80
+#82 := [unit-resolution #81 #52]: #69
+#59 := (= uf_1 0::int)
+#83 := (not #59)
+#89 := (iff #83 #50)
+#87 := (iff #59 #53)
+#88 := [commutativity]: #87
+#90 := [monotonicity #88]: #89
+#84 := (or #83 #30)
+#85 := [th-lemma]: #84
+#86 := [unit-resolution #85 #34]: #83
+#91 := [mp #86 #90]: #50
+#64 := -1::int
+#65 := (* -1::int #29)
+#66 := (+ uf_1 #65)
+#68 := (>= #66 0::int)
+#92 := (not #68)
+#93 := (or #92 #30)
+#94 := [th-lemma]: #93
+#95 := [unit-resolution #94 #34]: #92
+#96 := (or #57 #68)
+#97 := [th-lemma]: #96
+#98 := [unit-resolution #97 #95]: #57
+#76 := (or #58 #53 #54 #56)
+#77 := [def-axiom]: #76
+#99 := [unit-resolution #77 #98 #91 #82]: #54
+[unit-resolution #99 #112]: false
+unsat
+9791139ea2cfda88ae799477fc61e411aab42b18 673 0
#2 := false
#169 := 0::int
decl uf_2 :: int
@@ -4302,7 +4302,7 @@
#410 := [mp #349 #409]: #405
[unit-resolution #410 #710 #709 #697 #706]: false
unsat
-U7jSPEM53XYq3qs03aUczw 2291 0
+0d07e457c5602ba4a5dc70af32b6dc53a80a858c 2291 0
#2 := false
#6 := 0::int
decl z3name!0 :: int
@@ -6594,7 +6594,7 @@
#2323 := [unit-resolution #918 #2322]: #100
[unit-resolution #917 #2323 #2318]: false
unsat
-eqE7IAqFr0UIBuUsVDgHvw 52 0
+258b6cd4609a61b7800235c7f356739cfb8996c5 52 0
#2 := false
#8 := 1::real
decl uf_1 :: real
@@ -6647,7 +6647,7 @@
#29 := [asserted]: #13
[mp #29 #65]: false
unsat
-ADs4ZPiuUr7Xu7tk71NnEw 59 0
+3d1d0473f97c11d6c4d10f6e0313b2e2f4aac879 59 0
#2 := false
#55 := 0::int
#7 := 2::int
@@ -6707,7 +6707,7 @@
#144 := [unit-resolution #143 #28]: #58
[unit-resolution #144 #68]: false
unsat
-x2NmsblNl28xPXP2EG22rA 54 0
+f768cbe713eb8031e45b1a78d0f49a07f5398eb8 54 0
#2 := false
#5 := 2::int
decl uf_1 :: int
@@ -6762,7 +6762,7 @@
#139 := [unit-resolution #138 #27]: #127
[unit-resolution #139 #62]: false
unsat
-kfLiOGBz3RZx9wt+FS+hfg 118 0
+2c2bcacfbe018175dd39ce04dd5cbe02c800a0dd 118 0
#2 := false
#5 := 0::real
decl uf_1 :: real
@@ -6881,7 +6881,7 @@
#126 := [mp #101 #125]: #115
[unit-resolution #126 #132 #129]: false
unsat
-FPTJq9aN3ES4iIrHgaTv+A 208 0
+c8ed8eae868ac61c8f3a616f1b0b1df4032f4eac 208 0
#2 := false
#9 := 0::int
#11 := 4::int
@@ -7090,7 +7090,7 @@
#418 := [unit-resolution #417 #36]: #298
[th-lemma #418 #415 #412 #411 #99 #400 #397 #396]: false
unsat
-yN0aj3KferzvOSp2KlyNwg 24 0
+7beaddc803d2c23197634dc63d56d564292d85fe 24 0
#2 := false
#4 := (exists (vars (?x1 int)) false)
#5 := (not #4)
@@ -7115,7 +7115,7 @@
#22 := [asserted]: #6
[mp #22 #38]: false
unsat
-7iMPasu6AIeHm45slLCByA 24 0
+723fcd1ecb9fa59a7e0fede642f23063fb499818 24 0
#2 := false
#4 := (exists (vars (?x1 real)) false)
#5 := (not #4)
@@ -7140,13 +7140,13 @@
#22 := [asserted]: #6
[mp #22 #38]: false
unsat
-cv2pC2I0gIUYtVwtXngvXg 1 0
-unsat
-4r8/IxBBDH1ZqF0YfzLLTg 1 0
-unsat
-uj7n+C4nG462DNJy9Divrg 1 0
-unsat
-dn/LVwy1BXEOmtqdUBNhLw 73 0
+a72d0e977596e1fac0cccee600f0bf9d29ed71aa 1 0
+unsat
+70141a690f46561f859d3deed80b9611816f9f81 1 0
+unsat
+41b6ddffa2c7efc9285d0e0a65d74c4325ef6ddb 1 0
+unsat
+a08fcdd29520930b0a940df57c3d8266dbefd10f 73 0
#2 := false
#5 := 0::int
#8 := 1::int
@@ -7220,7 +7220,7 @@
#144 := [trans #142 #82]: #143
[mp #144 #146]: false
unsat
-VzZ1W5SEEis1AJp1qZz86g 82 0
+470993954e986ab72716000fd7da9fa600b05225 82 0
#2 := false
#5 := (:var 0 int)
#7 := 0::int
@@ -7303,7 +7303,7 @@
#30 := [asserted]: #14
[mp #30 #96]: false
unsat
-UoXgZh5LkmyNCmQEfEtnig 78 0
+40c93af1a084932780f95bda03b3df7712e01201 78 0
#2 := false
#5 := (:var 0 int)
#7 := 2::int
@@ -7382,7 +7382,7 @@
#31 := [asserted]: #15
[mp #31 #92]: false
unsat
-Qv4gVhCmOzC39uufV9ZpDA 61 0
+26b175ea54cef59293a917c6fb083751b00d312a 61 0
#2 := false
#9 := (:var 0 int)
#4 := 2::int
@@ -7444,7 +7444,7 @@
#30 := [asserted]: #14
[mp #30 #75]: false
unsat
-+j+tSj7aUImWej2XcTL9dw 111 0
+74037c10b4f126275ba21e7140b7f1e159b39ed9 111 0
#2 := false
#4 := 2::int
decl ?x1!1 :: int
@@ -7556,7 +7556,7 @@
#184 := [th-lemma]: #183
[unit-resolution #184 #127 #125 #126]: false
unsat
-kQRsBd9oowc7exsvsEgTUg 89 0
+628c1b88ca8fb09c896ae05059a52dc2f8e25db2 89 0
#2 := false
#4 := 0::int
decl ?x1!0 :: int
@@ -7646,7 +7646,7 @@
#167 := [unit-resolution #154 #90]: #166
[unit-resolution #167 #165 #162]: false
unsat
-VPjD8BtzPcTZKIRT4SA3Nw 83 2
+b7c4f9440c4594c46eee14ce57f17610bb7e2536 83 2
#2 := false
#5 := 0::int
#4 := (:var 0 int)
@@ -7732,7 +7732,7 @@
unsat
WARNING: failed to find a pattern for quantifier (quantifier id: k!2)
-DCV5zpDW3cC2A61VghqFkA 180 2
+7a9cc3ee85422788d981af84d181bd61d65f774c 180 2
#2 := false
#4 := 0::int
#5 := (:var 0 int)
@@ -7915,7 +7915,7 @@
unsat
WARNING: failed to find a pattern for quantifier (quantifier id: k!2)
-lYXJpXHB9nLXJbOsr9VH1w 68 0
+5201b12abd6b3d0f247a34c1fd9f443fc951c55f 68 0
#2 := false
#12 := 1::int
#9 := (:var 1 int)
@@ -7984,7 +7984,7 @@
#32 := [asserted]: #16
[mp #32 #83]: false
unsat
-jNvpOd8qnh73F8B6mQDrRw 107 0
+0f9091dc6853772b5280c29fc11ae1382022f24d 107 0
#2 := false
#4 := 0::int
decl ?x2!1 :: int
@@ -8092,7 +8092,7 @@
#123 := [and-elim #101]: #88
[th-lemma #123 #124 #125]: false
unsat
-QWWPBUGjgvTCpxqJ9oPGdQ 117 0
+a19e2cec45cb985989328595a0e06836a1e0fbc3 117 0
#2 := false
#4 := 0::int
decl ?x2!1 :: int
@@ -8210,7 +8210,7 @@
#188 := [unit-resolution #187 #110]: #98
[unit-resolution #188 #130]: false
unsat
-3r4MsKEvDJc1RWnNRxu/3Q 148 0
+34bf666106f50c4ee2e8834de4912d59c6e7d9d9 148 0
#2 := false
#144 := (not false)
#7 := 0::int
@@ -8359,7 +8359,7 @@
#158 := [mp #126 #157]: #153
[mp #158 #181]: false
unsat
-Q+cnHyqIFLGWsSlQkp3fEg 67 0
+1d6946d9384f22b76e98f04aff657c54e4fe51ad 67 0
#2 := false
#4 := (:var 0 int)
#5 := (pattern #4)
@@ -8427,9 +8427,9 @@
#30 := [asserted]: #14
[mp #30 #80]: false
unsat
-Q7HDzu4ER2dw+lHHM6YgFg 1 0
-unsat
-saejIG5KeeVxOolEIo3gtw 75 0
+d938f8b556e86b20a82e4661e3a61bad7d95357d 1 0
+unsat
+dfca84a72c9a54145743ea34eaa7c75e8665fd45 75 0
#2 := false
#6 := 1::int
decl uf_3 :: int
@@ -8505,7 +8505,7 @@
#32 := [asserted]: #16
[mp #32 #86]: false
unsat
-PPaoU5CzQFYr3LRpOsGPhQ 62 0
+2662a556257bfe403cd3fda75e9fe55964bc9dcd 62 0
#2 := false
decl uf_2 :: real
#6 := uf_2
@@ -8568,7 +8568,7 @@
#32 := [asserted]: #16
[mp #32 #74]: false
unsat
-hXKzem5+KYZMOj+GKxjszQ 141 0
+99d65e323e3d85dbdc2d8c52c169db4d46ead198 141 0
#2 := false
decl uf_4 :: int
#9 := uf_4
@@ -8710,7 +8710,7 @@
#45 := [asserted]: #29
[mp #45 #150]: false
unsat
-3D8WhjZTO7T824d7mwXcCA 252 0
+2e721ab2035f9845f1e87e78db6dfc67c28f6d40 252 0
#2 := false
#9 := 0::int
decl uf_2 :: (-> T1 int)
@@ -8963,7 +8963,7 @@
#539 := [unit-resolution #532 #451]: #557
[th-lemma #455 #539 #537 #546]: false
unsat
-kyphS4o71h68g2YhvYbQQQ 223 0
+5d4787d5f6bf7b62bda1a48bdd01dc6863801852 223 0
#2 := false
#23 := 3::int
decl uf_2 :: (-> T1 int)
@@ -9187,7 +9187,7 @@
#598 := [unit-resolution #593 #596]: #623
[th-lemma #152 #598 #139]: false
unsat
-M8P5WxpiY5AWxaJDBtXoLQ 367 0
+60689c41168db239dbf5f3a98d5f2bce0fef9e02 367 0
#2 := false
#9 := 0::int
decl uf_2 :: (-> T1 int)
@@ -9555,7 +9555,7 @@
#456 := [th-lemma]: #455
[unit-resolution #456 #464 #452]: false
unsat
-Xs4JZCKb5egkcPabsrodXg 302 0
+94b7ba760bb9dd467688fc28632e0ae8f6f51951 302 0
#2 := false
#9 := 0::int
decl uf_2 :: (-> T1 int)
@@ -9858,7 +9858,7 @@
#601 := [unit-resolution #615 #613]: #683
[th-lemma #623 #188 #601 #628]: false
unsat
-clMAi2WqMi360EjFURRGLg 458 0
+8d2fca14b1477934a0c7f4f6528bd3be029bba7b 458 0
#2 := false
#9 := 0::int
decl uf_2 :: (-> T1 int)
@@ -10317,7 +10317,7 @@
#350 := [unit-resolution #369 #367]: #368
[unit-resolution #350 #323]: false
unsat
-mu7O1os0t3tPqWZhwizjxw 161 0
+720080123967f7b12d5ac9ba2a5e5203400a16cd 161 0
#2 := false
#9 := 0::int
decl uf_3 :: int
@@ -10479,7 +10479,7 @@
#361 := [unit-resolution #639 #655]: #647
[th-lemma #656 #361 #261]: false
unsat
-08cmOtIT4NYs2PG/F3zeZw 557 0
+980ba021fe0853c3aa46377383e3f0f6fa9e2888 557 0
#2 := false
#9 := 0::int
decl uf_2 :: (-> T1 int)
@@ -11037,57 +11037,57 @@
#990 := [unit-resolution #501 #807]: #511
[unit-resolution #990 #989 #979]: false
unsat
-8HdmSMHHP2B8XMFzjNuw5Q 1 0
-unsat
-O4aM0+/isn2q5CrIefZjzg 1 0
-unsat
-t/ni9djl2DqxH0iKupZSwg 1 0
-unsat
-RumBGekdxZQaBF1HNa3x9w 1 0
-unsat
-Q9d+IbQ8chjKld71X6/zqw 1 0
-unsat
-PhC8zQV8hnJ6E2YYjZPGjQ 1 0
-unsat
-mieI2RhSp3bYaojlWH1A4A 1 0
-unsat
-pRSV6nBLconzrQz2zUrJ6g 1 0
-unsat
-Js0JfdwDoKq3YuilPPgeZw 1 0
-unsat
-GRIqjLUJiqXbo+pXhAeKIw 1 0
-unsat
-Bg5scsmPFp82+7Y2ScL6Wg 1 0
-unsat
-XD6zX6850dLxyfZSfNv30A 1 0
-unsat
-BG/HwJYnumvDICXxtBu/tA 1 0
-unsat
-YMc4t19sUMWbUkx3woxCmQ 1 0
-unsat
-YyD9IF72pKXGGKZTO7FY5Q 1 0
-unsat
-zRPsIUi+TEoz5fPWP0H9bQ 1 0
-unsat
-8ipTE8BOXpvSo/U6D4p3lA 1 0
-unsat
-MSzQywedZPsOE0CDxrrO0g 1 0
-unsat
-SryZuXv48ItET8NPIv07pA 1 0
-unsat
-qOMUQN18hYFl/wWt54lvbA 1 0
-unsat
-+njWXdn6fETK3/AjtiHjcA 1 0
-unsat
-5cQ7gJ33gzYTIIPA3hbBmQ 1 0
-unsat
-ZznT34cvumrP00mXZ3gcjw 1 0
-unsat
-//LQca1Et5RfhQJZA+CGCA 1 0
-unsat
-3ntxKz+kaQNfTrLzY9sVXw 1 0
-unsat
-4lL2Qo8ngE1EH1UdeN1Qng 43 0
+22877b17eafaba69b1f8a961a616fea28ae70d56 1 0
+unsat
+b5839159097bbd4e601a5681d1ca3493ec994a7c 1 0
+unsat
+90e1074350b5dcaae149781bcaa5d643b2ca9f48 1 0
+unsat
+08c7117fe974f5767051ed5aa61a27ce3084eb1d 1 0
+unsat
+858012417c9d327d8997f2a5dcb3da095ec65d34 1 0
+unsat
+84b2eee4890eaadb3638c7e522fb3237b3d476b0 1 0
+unsat
+8867717d9736308a2c27df0665a6e391b0562076 1 0
+unsat
+cd79c9a0488ab597d08dd9a0d6ac0f3647003bd6 1 0
+unsat
+395dd6c10b2a432137f9e3401cba8ec4dd64911b 1 0
+unsat
+17e3cc9534e04d86f095ec1a92c77d46d7dbb8e5 1 0
+unsat
+e046ea79beacf4bc3567b3b7f755232369d0c185 1 0
+unsat
+8ce4235464829d4be72e682f0c72bc5e3c6902d0 1 0
+unsat
+656a40b977d7716264443900d6bdb4d3d117d52f 1 0
+unsat
+ec27a57e58719625ff71dd4d68ed53a3851efb5c 1 0
+unsat
+2c3c366b8488ca0991cc767c94cdb78b18db9d5f 1 0
+unsat
+5894f6f19250b12885e38f54eae81f143b58fa01 1 0
+unsat
+e150815d9eb1ec168805b5501d7f4b2e378dd883 1 0
+unsat
+396d6254e993f414335de9378150e486d3cfcd4e 1 0
+unsat
+96014c61f582a811aff25ad7fa62b575b830fa8b 1 0
+unsat
+10580b87c0d062c9854e79d16047a53d045ccfac 1 0
+unsat
+87b5f388df1f43cc02ac0fa0d6944eb8cd8f8f50 1 0
+unsat
+92a5e2bb68f74b9e7dd3a9ef79ea641e9700d563 1 0
+unsat
+152e0f0f0a04b399b057beae92ae0455408b224f 1 0
+unsat
+41925af4711748a6e411453f2465920a1c6ffb8e 1 0
+unsat
+757462716f4a2619a1410bdca3faa2d058042c10 1 0
+unsat
+abdeeb4668a63f19473d6da94232379344d99fea 43 0
#2 := false
#6 := 0::int
decl uf_1 :: (-> bv[2] int)
@@ -11131,9 +11131,9 @@
#287 := [th-lemma]: #627
[unit-resolution #287 #47 #635]: false
unsat
-+xe3O927LrflFUE6NDqRlA 1 0
-unsat
-JPoL7fPYhqhAkjUiVF+THQ 50 0
+2e4bd6e1ab7bfedbefd3ada8caf5332ba5651065 1 0
+unsat
+96d2e1aad559535e781a07ee5e55bb19caef769c 50 0
#2 := false
decl uf_6 :: T2
#23 := uf_6
@@ -11184,7 +11184,7 @@
#66 := [asserted]: #26
[unit-resolution #66 #235]: false
unsat
-l23ZDmd0VbO/Q+uO5EtabA 105 0
+008ad82b0a62ff600e66e8f2cc72b5795c7c5032 105 0
#2 := false
decl uf_6 :: (-> T4 T2)
decl uf_10 :: T4
@@ -11290,7 +11290,7 @@
#110 := [asserted]: #46
[unit-resolution #110 #238]: false
unsat
-GZjffeZPQnL3OyLCvxdCpg 181 0
+0a15a60660bc0c3f06688ad5de50ed50a24d0df0 181 0
#2 := false
decl uf_1 :: (-> T1 T2 T3)
decl uf_3 :: T2
@@ -11472,7 +11472,7 @@
#76 := [asserted]: #38
[unit-resolution #76 #489]: false
unsat
-i6jCzzRosHYE0w7sF1Nraw 62 0
+dde2364f13fc9ce8af00c7a02bfea9a6c979a805 62 0
#2 := false
decl up_4 :: (-> T1 T2 bool)
decl uf_3 :: T2
@@ -11535,7 +11535,7 @@
#73 := [unit-resolution #71 #68]: #72
[unit-resolution #73 #59 #61]: false
unsat
-YZHSyhN2TGlpe+vpkzWrgQ 115 0
+3cb6787c7063b0a04bff9e2bfa4203b73903dabe 115 0
#2 := false
decl up_2 :: (-> T2 bool)
decl uf_3 :: T2
@@ -11651,7 +11651,7 @@
#560 := [mp #344 #559]: #557
[unit-resolution #560 #576 #561]: false
unsat
-TibRlXkU+X+1+zGPYTiT0g 464 0
+dcb635ada6f2482ea11de30713fe962bcb57c9ad 464 0
#2 := false
decl uf_2 :: (-> T2 T3 T3)
decl uf_4 :: T3
@@ -12116,7 +12116,7 @@
#177 := [asserted]: #53
[unit-resolution #177 #793]: false
unsat
-DJPKxi9AO25zGBcs5kxUrw 21 0
+ca467a37d809de06757809cab1cd30e08586674f 21 0
#2 := false
decl up_1 :: (-> T1 bool)
#4 := (:var 0 T1)
@@ -12138,7 +12138,7 @@
#25 := [asserted]: #9
[mp #25 #34]: false
unsat
-i5PnMbuM9mWv5LnVszz9+g 366 0
+f6e9645ebbb8739d12b8e96bc83ddef33c7d53bc 366 0
#2 := false
decl uf_1 :: (-> int T1)
#37 := 6::int
@@ -12505,7 +12505,7 @@
#182 := [asserted]: #40
[unit-resolution #182 #399]: false
unsat
-K2SXMHU6QCZJ8TRs6zjKRg 408 0
+b15b527236338a437e355afcde07dd3c6dfc451f 408 0
#2 := false
#22 := 0::int
#8 := 2::int
@@ -12914,7 +12914,7 @@
#375 := [unit-resolution #374 #407]: #708
[th-lemma #375 #370 #465]: false
unsat
-1DhSL9G2fGRGmuI8IaMNOA 50 0
+71b79533f46a8514b2fc189fc382867d50f52a9a 50 0
#2 := false
decl up_35 :: (-> int bool)
#112 := 1::int
@@ -12965,7 +12965,7 @@
#504 := [quant-inst]: #503
[unit-resolution #504 #916 #297]: false
unsat
-dyXROdcPFSd36N3K7dpmDw 506 0
+e40ba7aed8ae5409337a331038da0587c03354d6 506 0
#2 := false
decl uf_17 :: (-> T8 T3)
decl uf_18 :: (-> T1 T8)
--- a/src/HOL/SMT/SMT_Base.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOL/SMT/SMT_Base.thy Thu Mar 25 17:56:31 2010 +0100
@@ -8,7 +8,7 @@
imports Real "~~/src/HOL/Word/Word"
"~~/src/HOL/Decision_Procs/Dense_Linear_Order"
uses
- "~~/src/Tools/Cache_IO/cache_io.ML"
+ "~~/src/Tools/cache_io.ML"
("Tools/smt_normalize.ML")
("Tools/smt_monomorph.ML")
("Tools/smt_translate.ML")
--- a/src/HOL/Tools/Datatype/datatype_codegen.ML Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOL/Tools/Datatype/datatype_codegen.ML Thu Mar 25 17:56:31 2010 +0100
@@ -358,7 +358,7 @@
[trueprop $ false_eq (t1, t2), trueprop $ false_eq (t2, t1)];
val distincts = maps prep_distinct (snd (nth (Datatype_Prop.make_distincts [descr] vs) index));
val refl = HOLogic.mk_Trueprop (true_eq (Free ("x", ty), Free ("x", ty)));
- val simpset = Simplifier.context (ProofContext.init thy) (HOL_basic_ss addsimps
+ val simpset = Simplifier.global_context thy (HOL_basic_ss addsimps
(map Simpdata.mk_eq (@{thm eq} :: @{thm eq_True} :: inject_thms @ distinct_thms)));
fun prove prop = Skip_Proof.prove_global thy [] [] prop (K (ALLGOALS (simp_tac simpset)))
|> Simpdata.mk_eq;
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Thu Mar 25 17:56:31 2010 +0100
@@ -177,10 +177,10 @@
val new_defs = map new_def assms
val (definition, thy') = thy
|> Sign.add_consts_i [(Binding.name constname, constT, NoSyn)]
- |> fold_map Drule.add_axiom (map_index
- (fn (i, t) => (Binding.name (constname ^ "_def" ^ string_of_int i), t)) new_defs)
+ |> fold_map Specification.axiom (map_index
+ (fn (i, t) => ((Binding.name (constname ^ "_def" ^ string_of_int i), []), t)) new_defs)
in
- (lhs, ((full_constname, definition) :: defs, thy'))
+ (lhs, ((full_constname, map Drule.export_without_context definition) :: defs, thy'))
end
else
(atom, (defs, thy)))
--- a/src/HOL/Tools/choice_specification.ML Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOL/Tools/choice_specification.ML Thu Mar 25 17:56:31 2010 +0100
@@ -44,9 +44,10 @@
let
fun process [] (thy,tm) =
let
- val (thm, thy') = Drule.add_axiom (Binding.name axname, HOLogic.mk_Trueprop tm) thy
+ val (thm, thy') =
+ Specification.axiom ((Binding.name axname, []), HOLogic.mk_Trueprop tm) thy
in
- (thy', thm)
+ (thy', Drule.export_without_context thm)
end
| process ((thname,cname,covld)::cos) (thy,tm) =
case tm of
--- a/src/HOL/ex/Predicate_Compile_Alternative_Defs.thy Thu Mar 25 17:55:55 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,107 +0,0 @@
-theory Predicate_Compile_Alternative_Defs
-imports "../Predicate_Compile"
-begin
-
-section {* Common constants *}
-
-declare HOL.if_bool_eq_disj[code_pred_inline]
-
-setup {* Predicate_Compile_Data.ignore_consts [@{const_name Let}] *}
-
-section {* Pairs *}
-
-setup {* Predicate_Compile_Data.ignore_consts [@{const_name fst}, @{const_name snd}, @{const_name split}] *}
-
-section {* Bounded quantifiers *}
-
-declare Ball_def[code_pred_inline]
-declare Bex_def[code_pred_inline]
-
-section {* Set operations *}
-
-declare Collect_def[code_pred_inline]
-declare mem_def[code_pred_inline]
-
-declare eq_reflection[OF empty_def, code_pred_inline]
-declare insert_code[code_pred_def]
-
-declare subset_iff[code_pred_inline]
-
-declare Int_def[code_pred_inline]
-declare eq_reflection[OF Un_def, code_pred_inline]
-declare eq_reflection[OF UNION_def, code_pred_inline]
-
-lemma Diff[code_pred_inline]:
- "(A - B) = (%x. A x \<and> \<not> B x)"
-by (auto simp add: mem_def)
-
-lemma set_equality[code_pred_inline]:
- "(A = B) = ((\<forall>x. A x \<longrightarrow> B x) \<and> (\<forall>x. B x \<longrightarrow> A x))"
-by (fastsimp simp add: mem_def)
-
-section {* Setup for Numerals *}
-
-setup {* Predicate_Compile_Data.ignore_consts [@{const_name number_of}] *}
-setup {* Predicate_Compile_Data.keep_functions [@{const_name number_of}] *}
-
-setup {* Predicate_Compile_Data.ignore_consts [@{const_name div}, @{const_name mod}, @{const_name times}] *}
-
-section {* Alternative list definitions *}
-
-text {* size simps are not yet added to the Spec_Rules interface. So they are just added manually here! *}
-
-lemma [code_pred_def]:
- "length [] = 0"
- "length (x # xs) = Suc (length xs)"
-by auto
-
-subsection {* Alternative rules for set *}
-
-lemma set_ConsI1 [code_pred_intro]:
- "set (x # xs) x"
-unfolding mem_def[symmetric, of _ x]
-by auto
-
-lemma set_ConsI2 [code_pred_intro]:
- "set xs x ==> set (x' # xs) x"
-unfolding mem_def[symmetric, of _ x]
-by auto
-
-code_pred [skip_proof] set
-proof -
- case set
- from this show thesis
- apply (case_tac xb)
- apply auto
- unfolding mem_def[symmetric, of _ xc]
- apply auto
- unfolding mem_def
- apply fastsimp
- done
-qed
-
-subsection {* Alternative rules for list_all2 *}
-
-lemma list_all2_NilI [code_pred_intro]: "list_all2 P [] []"
-by auto
-
-lemma list_all2_ConsI [code_pred_intro]: "list_all2 P xs ys ==> P x y ==> list_all2 P (x#xs) (y#ys)"
-by auto
-
-code_pred [skip_proof] list_all2
-proof -
- case list_all2
- from this show thesis
- apply -
- apply (case_tac xb)
- apply (case_tac xc)
- apply auto
- apply (case_tac xc)
- apply auto
- apply fastsimp
- done
-qed
-
-
-
-end
\ No newline at end of file
--- a/src/HOL/ex/Predicate_Compile_Quickcheck.thy Thu Mar 25 17:55:55 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,39 +0,0 @@
-(* Author: Lukas Bulwahn, TU Muenchen *)
-
-header {* A Prototype of Quickcheck based on the Predicate Compiler *}
-
-theory Predicate_Compile_Quickcheck
-imports "../Predicate_Compile" Quickcheck Predicate_Compile_Alternative_Defs
-uses "../Tools/Predicate_Compile/predicate_compile_quickcheck.ML"
-begin
-
-setup {* Quickcheck.add_generator ("predicate_compile_wo_ff", Predicate_Compile_Quickcheck.quickcheck_compile_term false true 8) *}
-setup {* Quickcheck.add_generator ("predicate_compile_ff_fs", Predicate_Compile_Quickcheck.quickcheck_compile_term true true 8) *}
-setup {* Quickcheck.add_generator ("predicate_compile_ff_nofs", Predicate_Compile_Quickcheck.quickcheck_compile_term true false 8) *}
-
-(*
-datatype alphabet = a | b
-
-inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
- "[] \<in> S\<^isub>1"
-| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
-| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
-| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
-| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
-| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
-
-ML {* set Toplevel.debug *}
-
-declare mem_def[code_pred_inline] Collect_def[code_pred_inline]
-
-lemma
- "w \<in> S\<^isub>1p \<Longrightarrow> w = []"
-quickcheck[generator = predicate_compile, iterations=1]
-oops
-
-theorem S\<^isub>1_sound:
-"w \<in> S\<^isub>1p \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-quickcheck[generator=predicate_compile, size=15]
-oops
-*)
-end
\ No newline at end of file
--- a/src/HOL/ex/Predicate_Compile_Quickcheck_ex.thy Thu Mar 25 17:55:55 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,337 +0,0 @@
-theory Predicate_Compile_Quickcheck_ex
-imports Predicate_Compile_Quickcheck
-begin
-
-section {* Sets *}
-
-lemma "x \<in> {(1::nat)} ==> False"
-quickcheck[generator=predicate_compile_wo_ff, iterations=10]
-oops
-
-lemma "x \<in> {Suc 0, Suc (Suc 0)} ==> x \<noteq> Suc 0"
-quickcheck[generator=predicate_compile_wo_ff]
-oops
-
-lemma "x \<in> {Suc 0, Suc (Suc 0)} ==> x = Suc 0"
-quickcheck[generator=predicate_compile_wo_ff]
-oops
-
-lemma "x \<in> {Suc 0, Suc (Suc 0)} ==> x <= Suc 0"
-quickcheck[generator=predicate_compile_wo_ff]
-oops
-
-section {* Numerals *}
-
-lemma
- "x \<in> {1, 2, (3::nat)} ==> x = 1 \<or> x = 2"
-quickcheck[generator=predicate_compile_wo_ff]
-oops
-
-lemma "x \<in> {1, 2, (3::nat)} ==> x < 3"
-quickcheck[generator=predicate_compile_wo_ff]
-oops
-
-lemma
- "x \<in> {1, 2} \<union> {3, 4} ==> x = (1::nat) \<or> x = (2::nat)"
-quickcheck[generator=predicate_compile_wo_ff]
-oops
-
-section {* Context Free Grammar *}
-
-datatype alphabet = a | b
-
-inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
- "[] \<in> S\<^isub>1"
-| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
-| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
-| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
-| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
-| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
-
-lemma
- "w \<in> S\<^isub>1 \<Longrightarrow> w = []"
-quickcheck[generator = predicate_compile_ff_nofs, iterations=1]
-oops
-
-theorem S\<^isub>1_sound:
-"w \<in> S\<^isub>1 \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-quickcheck[generator=predicate_compile_ff_nofs, size=15]
-oops
-
-
-inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
- "[] \<in> S\<^isub>2"
-| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
-| "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2"
-| "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
-| "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
-| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
-(*
-code_pred [random_dseq inductify] S\<^isub>2 .
-thm S\<^isub>2.random_dseq_equation
-thm A\<^isub>2.random_dseq_equation
-thm B\<^isub>2.random_dseq_equation
-
-values [random_dseq 1, 2, 8] 10 "{x. S\<^isub>2 x}"
-
-lemma "w \<in> S\<^isub>2 ==> w \<noteq> [] ==> w \<noteq> [b, a] ==> w \<in> {}"
-quickcheck[generator=predicate_compile, size=8]
-oops
-
-lemma "[x <- w. x = a] = []"
-quickcheck[generator=predicate_compile]
-oops
-
-declare list.size(3,4)[code_pred_def]
-
-(*
-lemma "length ([x \<leftarrow> w. x = a]) = (0::nat)"
-quickcheck[generator=predicate_compile]
-oops
-*)
-
-lemma
-"w \<in> S\<^isub>2 ==> length [x \<leftarrow> w. x = a] <= Suc (Suc 0)"
-quickcheck[generator=predicate_compile, size = 10, iterations = 1]
-oops
-*)
-theorem S\<^isub>2_sound:
-"w \<in> S\<^isub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-quickcheck[generator=predicate_compile_ff_nofs, size=5, iterations=10]
-oops
-
-inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
- "[] \<in> S\<^isub>3"
-| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3"
-| "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3"
-| "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3"
-| "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3"
-| "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3"
-
-code_pred [inductify] S\<^isub>3 .
-thm S\<^isub>3.equation
-(*
-values 10 "{x. S\<^isub>3 x}"
-*)
-
-
-lemma S\<^isub>3_sound:
-"w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-quickcheck[generator=predicate_compile_ff_fs, size=10, iterations=10]
-oops
-
-lemma "\<not> (length w > 2) \<or> \<not> (length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b])"
-quickcheck[size=10, generator = predicate_compile_ff_fs]
-oops
-
-theorem S\<^isub>3_complete:
-"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. b = x] \<longrightarrow> w \<in> S\<^isub>3"
-(*quickcheck[generator=SML]*)
-quickcheck[generator=predicate_compile_ff_fs, size=10, iterations=100]
-oops
-
-
-inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
- "[] \<in> S\<^isub>4"
-| "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4"
-| "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4"
-| "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4"
-| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
-| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
-| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
-
-theorem S\<^isub>4_sound:
-"w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-quickcheck[generator = predicate_compile_ff_nofs, size=5, iterations=1]
-oops
-
-theorem S\<^isub>4_complete:
-"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4"
-quickcheck[generator = predicate_compile_ff_nofs, size=5, iterations=1]
-oops
-
-hide const a b
-
-subsection {* Lexicographic order *}
-(* TODO *)
-(*
-lemma
- "(u, v) : lexord r ==> (x @ u, y @ v) : lexord r"
-oops
-*)
-subsection {* IMP *}
-
-types
- var = nat
- state = "int list"
-
-datatype com =
- Skip |
- Ass var "int" |
- Seq com com |
- IF "state list" com com |
- While "state list" com
-
-inductive exec :: "com => state => state => bool" where
- "exec Skip s s" |
- "exec (Ass x e) s (s[x := e])" |
- "exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" |
- "s \<in> set b ==> exec c1 s t ==> exec (IF b c1 c2) s t" |
- "s \<notin> set b ==> exec c2 s t ==> exec (IF b c1 c2) s t" |
- "s \<notin> set b ==> exec (While b c) s s" |
- "s1 \<in> set b ==> exec c s1 s2 ==> exec (While b c) s2 s3 ==> exec (While b c) s1 s3"
-
-code_pred [random_dseq] exec .
-
-values [random_dseq 1, 2, 3] 10 "{(c, s, s'). exec c s s'}"
-
-lemma
- "exec c s s' ==> exec (Seq c c) s s'"
-(*quickcheck[generator = predicate_compile_wo_ff, size=2, iterations=10]*)
-oops
-
-subsection {* Lambda *}
-
-datatype type =
- Atom nat
- | Fun type type (infixr "\<Rightarrow>" 200)
-
-datatype dB =
- Var nat
- | App dB dB (infixl "\<degree>" 200)
- | Abs type dB
-
-primrec
- nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
-where
- "[]\<langle>i\<rangle> = None"
-| "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
-
-inductive nth_el' :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool"
-where
- "nth_el' (x # xs) 0 x"
-| "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y"
-
-inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" ("_ \<turnstile> _ : _" [50, 50, 50] 50)
- where
- Var [intro!]: "nth_el' env x T \<Longrightarrow> env \<turnstile> Var x : T"
- | Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)"
- | App [intro!]: "env \<turnstile> s : U \<Rightarrow> T \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U"
-
-primrec
- lift :: "[dB, nat] => dB"
-where
- "lift (Var i) k = (if i < k then Var i else Var (i + 1))"
- | "lift (s \<degree> t) k = lift s k \<degree> lift t k"
- | "lift (Abs T s) k = Abs T (lift s (k + 1))"
-
-primrec
- subst :: "[dB, dB, nat] => dB" ("_[_'/_]" [300, 0, 0] 300)
-where
- subst_Var: "(Var i)[s/k] =
- (if k < i then Var (i - 1) else if i = k then s else Var i)"
- | subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
- | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"
-
-inductive beta :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<beta>" 50)
- where
- beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
- | appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
- | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
- | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"
-
-lemma
- "\<Gamma> \<turnstile> t : U \<Longrightarrow> t \<rightarrow>\<^sub>\<beta> t' \<Longrightarrow> \<Gamma> \<turnstile> t' : U"
-quickcheck[generator = predicate_compile_ff_fs, size = 7, iterations = 10]
-oops
-
-subsection {* JAD *}
-
-definition matrix :: "('a :: semiring_0) list list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
- "matrix M rs cs \<longleftrightarrow> (\<forall> row \<in> set M. length row = cs) \<and> length M = rs"
-(*
-code_pred [random_dseq inductify] matrix .
-thm matrix.random_dseq_equation
-
-thm matrix_aux.random_dseq_equation
-
-values [random_dseq 3, 2] 10 "{(M, rs, cs). matrix (M:: int list list) rs cs}"
-*)
-lemma [code_pred_intro]:
- "matrix [] 0 m"
- "matrix xss n m ==> length xs = m ==> matrix (xs # xss) (Suc n) m"
-proof -
- show "matrix [] 0 m" unfolding matrix_def by auto
-next
- show "matrix xss n m ==> length xs = m ==> matrix (xs # xss) (Suc n) m"
- unfolding matrix_def by auto
-qed
-
-code_pred [random_dseq inductify] matrix
- apply (cases x)
- unfolding matrix_def apply fastsimp
- apply fastsimp done
-
-
-values [random_dseq 2, 2, 15] 6 "{(M::int list list, n, m). matrix M n m}"
-
-definition "scalar_product v w = (\<Sum> (x, y)\<leftarrow>zip v w. x * y)"
-
-definition mv :: "('a \<Colon> semiring_0) list list \<Rightarrow> 'a list \<Rightarrow> 'a list"
- where [simp]: "mv M v = map (scalar_product v) M"
-text {*
- This defines the matrix vector multiplication. To work properly @{term
-"matrix M m n \<and> length v = n"} must hold.
-*}
-
-subsection "Compressed matrix"
-
-definition "sparsify xs = [i \<leftarrow> zip [0..<length xs] xs. snd i \<noteq> 0]"
-(*
-lemma sparsify_length: "(i, x) \<in> set (sparsify xs) \<Longrightarrow> i < length xs"
- by (auto simp: sparsify_def set_zip)
-
-lemma listsum_sparsify[simp]:
- fixes v :: "('a \<Colon> semiring_0) list"
- assumes "length w = length v"
- shows "(\<Sum>x\<leftarrow>sparsify w. (\<lambda>(i, x). v ! i) x * snd x) = scalar_product v w"
- (is "(\<Sum>x\<leftarrow>_. ?f x) = _")
- unfolding sparsify_def scalar_product_def
- using assms listsum_map_filter[where f="?f" and P="\<lambda> i. snd i \<noteq> (0::'a)"]
- by (simp add: listsum_setsum)
-*)
-definition [simp]: "unzip w = (map fst w, map snd w)"
-
-primrec insert :: "('a \<Rightarrow> 'b \<Colon> linorder) => 'a \<Rightarrow> 'a list => 'a list" where
- "insert f x [] = [x]" |
- "insert f x (y # ys) = (if f y < f x then y # insert f x ys else x # y # ys)"
-
-primrec sort :: "('a \<Rightarrow> 'b \<Colon> linorder) \<Rightarrow> 'a list => 'a list" where
- "sort f [] = []" |
- "sort f (x # xs) = insert f x (sort f xs)"
-
-definition
- "length_permutate M = (unzip o sort (length o snd)) (zip [0 ..< length M] M)"
-(*
-definition
- "transpose M = [map (\<lambda> xs. xs ! i) (takeWhile (\<lambda> xs. i < length xs) M). i \<leftarrow> [0 ..< length (M ! 0)]]"
-*)
-definition
- "inflate upds = foldr (\<lambda> (i, x) upds. upds[i := x]) upds (replicate (length upds) 0)"
-
-definition
- "jad = apsnd transpose o length_permutate o map sparsify"
-
-definition
- "jad_mv v = inflate o split zip o apsnd (map listsum o transpose o map (map (\<lambda> (i, x). v ! i * x)))"
-
-lemma "matrix (M::int list list) rs cs \<Longrightarrow> False"
-quickcheck[generator = predicate_compile_ff_nofs, size = 6]
-oops
-
-lemma
- "\<lbrakk> matrix M rs cs ; length v = cs \<rbrakk> \<Longrightarrow> jad_mv v (jad M) = mv M v"
-quickcheck[generator = predicate_compile_wo_ff]
-oops
-
-end
\ No newline at end of file
--- a/src/HOL/ex/Predicate_Compile_ex.thy Thu Mar 25 17:55:55 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1278 +0,0 @@
-theory Predicate_Compile_ex
-imports Predicate_Compile_Alternative_Defs
-begin
-
-subsection {* Basic predicates *}
-
-inductive False' :: "bool"
-
-code_pred (expected_modes: bool) False' .
-code_pred [dseq] False' .
-code_pred [random_dseq] False' .
-
-values [expected "{}" pred] "{x. False'}"
-values [expected "{}" dseq 1] "{x. False'}"
-values [expected "{}" random_dseq 1, 1, 1] "{x. False'}"
-
-value "False'"
-
-
-inductive True' :: "bool"
-where
- "True ==> True'"
-
-code_pred True' .
-code_pred [dseq] True' .
-code_pred [random_dseq] True' .
-
-thm True'.equation
-thm True'.dseq_equation
-thm True'.random_dseq_equation
-values [expected "{()}" ]"{x. True'}"
-values [expected "{}" dseq 0] "{x. True'}"
-values [expected "{()}" dseq 1] "{x. True'}"
-values [expected "{()}" dseq 2] "{x. True'}"
-values [expected "{}" random_dseq 1, 1, 0] "{x. True'}"
-values [expected "{}" random_dseq 1, 1, 1] "{x. True'}"
-values [expected "{()}" random_dseq 1, 1, 2] "{x. True'}"
-values [expected "{()}" random_dseq 1, 1, 3] "{x. True'}"
-
-inductive EmptySet :: "'a \<Rightarrow> bool"
-
-code_pred (expected_modes: o => bool, i => bool) EmptySet .
-
-definition EmptySet' :: "'a \<Rightarrow> bool"
-where "EmptySet' = {}"
-
-code_pred (expected_modes: o => bool, i => bool) [inductify] EmptySet' .
-
-inductive EmptyRel :: "'a \<Rightarrow> 'b \<Rightarrow> bool"
-
-code_pred (expected_modes: o => o => bool, i => o => bool, o => i => bool, i => i => bool) EmptyRel .
-
-inductive EmptyClosure :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
-for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
-
-code_pred
- (expected_modes: (o => o => bool) => o => o => bool, (o => o => bool) => i => o => bool,
- (o => o => bool) => o => i => bool, (o => o => bool) => i => i => bool,
- (i => o => bool) => o => o => bool, (i => o => bool) => i => o => bool,
- (i => o => bool) => o => i => bool, (i => o => bool) => i => i => bool,
- (o => i => bool) => o => o => bool, (o => i => bool) => i => o => bool,
- (o => i => bool) => o => i => bool, (o => i => bool) => i => i => bool,
- (i => i => bool) => o => o => bool, (i => i => bool) => i => o => bool,
- (i => i => bool) => o => i => bool, (i => i => bool) => i => i => bool)
- EmptyClosure .
-
-thm EmptyClosure.equation
-
-(* TODO: inductive package is broken!
-inductive False'' :: "bool"
-where
- "False \<Longrightarrow> False''"
-
-code_pred (expected_modes: []) False'' .
-
-inductive EmptySet'' :: "'a \<Rightarrow> bool"
-where
- "False \<Longrightarrow> EmptySet'' x"
-
-code_pred (expected_modes: [1]) EmptySet'' .
-code_pred (expected_modes: [], [1]) [inductify] EmptySet'' .
-*)
-
-consts a' :: 'a
-
-inductive Fact :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
-where
-"Fact a' a'"
-
-code_pred (expected_modes: o => o => bool, i => o => bool, o => i => bool, i => i => bool) Fact .
-
-inductive zerozero :: "nat * nat => bool"
-where
- "zerozero (0, 0)"
-
-code_pred (expected_modes: i => bool, i * o => bool, o * i => bool, o => bool) zerozero .
-code_pred [dseq] zerozero .
-code_pred [random_dseq] zerozero .
-
-thm zerozero.equation
-thm zerozero.dseq_equation
-thm zerozero.random_dseq_equation
-
-text {* We expect the user to expand the tuples in the values command.
-The following values command is not supported. *}
-(*values "{x. zerozero x}" *)
-text {* Instead, the user must type *}
-values "{(x, y). zerozero (x, y)}"
-
-values [expected "{}" dseq 0] "{(x, y). zerozero (x, y)}"
-values [expected "{(0::nat, 0::nat)}" dseq 1] "{(x, y). zerozero (x, y)}"
-values [expected "{(0::nat, 0::nat)}" dseq 2] "{(x, y). zerozero (x, y)}"
-values [expected "{}" random_dseq 1, 1, 2] "{(x, y). zerozero (x, y)}"
-values [expected "{(0::nat, 0:: nat)}" random_dseq 1, 1, 3] "{(x, y). zerozero (x, y)}"
-
-inductive nested_tuples :: "((int * int) * int * int) => bool"
-where
- "nested_tuples ((0, 1), 2, 3)"
-
-code_pred nested_tuples .
-
-inductive JamesBond :: "nat => int => code_numeral => bool"
-where
- "JamesBond 0 0 7"
-
-code_pred JamesBond .
-
-values [expected "{(0::nat, 0::int , 7::code_numeral)}"] "{(a, b, c). JamesBond a b c}"
-values [expected "{(0::nat, 7::code_numeral, 0:: int)}"] "{(a, c, b). JamesBond a b c}"
-values [expected "{(0::int, 0::nat, 7::code_numeral)}"] "{(b, a, c). JamesBond a b c}"
-values [expected "{(0::int, 7::code_numeral, 0::nat)}"] "{(b, c, a). JamesBond a b c}"
-values [expected "{(7::code_numeral, 0::nat, 0::int)}"] "{(c, a, b). JamesBond a b c}"
-values [expected "{(7::code_numeral, 0::int, 0::nat)}"] "{(c, b, a). JamesBond a b c}"
-
-values [expected "{(7::code_numeral, 0::int)}"] "{(a, b). JamesBond 0 b a}"
-values [expected "{(7::code_numeral, 0::nat)}"] "{(c, a). JamesBond a 0 c}"
-values [expected "{(0::nat, 7::code_numeral)}"] "{(a, c). JamesBond a 0 c}"
-
-
-subsection {* Alternative Rules *}
-
-datatype char = C | D | E | F | G | H
-
-inductive is_C_or_D
-where
- "(x = C) \<or> (x = D) ==> is_C_or_D x"
-
-code_pred (expected_modes: i => bool) is_C_or_D .
-thm is_C_or_D.equation
-
-inductive is_D_or_E
-where
- "(x = D) \<or> (x = E) ==> is_D_or_E x"
-
-lemma [code_pred_intro]:
- "is_D_or_E D"
-by (auto intro: is_D_or_E.intros)
-
-lemma [code_pred_intro]:
- "is_D_or_E E"
-by (auto intro: is_D_or_E.intros)
-
-code_pred (expected_modes: o => bool, i => bool) is_D_or_E
-proof -
- case is_D_or_E
- from this(1) show thesis
- proof
- fix xa
- assume x: "x = xa"
- assume "xa = D \<or> xa = E"
- from this show thesis
- proof
- assume "xa = D" from this x is_D_or_E(2) show thesis by simp
- next
- assume "xa = E" from this x is_D_or_E(3) show thesis by simp
- qed
- qed
-qed
-
-thm is_D_or_E.equation
-
-inductive is_F_or_G
-where
- "x = F \<or> x = G ==> is_F_or_G x"
-
-lemma [code_pred_intro]:
- "is_F_or_G F"
-by (auto intro: is_F_or_G.intros)
-
-lemma [code_pred_intro]:
- "is_F_or_G G"
-by (auto intro: is_F_or_G.intros)
-
-inductive is_FGH
-where
- "is_F_or_G x ==> is_FGH x"
-| "is_FGH H"
-
-text {* Compilation of is_FGH requires elimination rule for is_F_or_G *}
-
-code_pred (expected_modes: o => bool, i => bool) is_FGH
-proof -
- case is_F_or_G
- from this(1) show thesis
- proof
- fix xa
- assume x: "x = xa"
- assume "xa = F \<or> xa = G"
- from this show thesis
- proof
- assume "xa = F"
- from this x is_F_or_G(2) show thesis by simp
- next
- assume "xa = G"
- from this x is_F_or_G(3) show thesis by simp
- qed
- qed
-qed
-
-subsection {* Preprocessor Inlining *}
-
-definition "equals == (op =)"
-
-inductive zerozero' :: "nat * nat => bool" where
- "equals (x, y) (0, 0) ==> zerozero' (x, y)"
-
-code_pred (expected_modes: i => bool) zerozero' .
-
-lemma zerozero'_eq: "zerozero' x == zerozero x"
-proof -
- have "zerozero' = zerozero"
- apply (auto simp add: mem_def)
- apply (cases rule: zerozero'.cases)
- apply (auto simp add: equals_def intro: zerozero.intros)
- apply (cases rule: zerozero.cases)
- apply (auto simp add: equals_def intro: zerozero'.intros)
- done
- from this show "zerozero' x == zerozero x" by auto
-qed
-
-declare zerozero'_eq [code_pred_inline]
-
-definition "zerozero'' x == zerozero' x"
-
-text {* if preprocessing fails, zerozero'' will not have all modes. *}
-
-code_pred (expected_modes: i * i => bool, i * o => bool, o * i => bool, o => bool) [inductify] zerozero'' .
-
-subsection {* Sets and Numerals *}
-
-definition
- "one_or_two = {Suc 0, (Suc (Suc 0))}"
-
-code_pred [inductify] one_or_two .
-
-code_pred [dseq] one_or_two .
-code_pred [random_dseq] one_or_two .
-thm one_or_two.dseq_equation
-values [expected "{Suc 0::nat, 2::nat}"] "{x. one_or_two x}"
-values [random_dseq 0,0,10] 3 "{x. one_or_two x}"
-
-inductive one_or_two' :: "nat => bool"
-where
- "one_or_two' 1"
-| "one_or_two' 2"
-
-code_pred one_or_two' .
-thm one_or_two'.equation
-
-values "{x. one_or_two' x}"
-
-definition one_or_two'':
- "one_or_two'' == {1, (2::nat)}"
-
-code_pred [inductify] one_or_two'' .
-thm one_or_two''.equation
-
-values "{x. one_or_two'' x}"
-
-subsection {* even predicate *}
-
-inductive even :: "nat \<Rightarrow> bool" and odd :: "nat \<Rightarrow> bool" where
- "even 0"
- | "even n \<Longrightarrow> odd (Suc n)"
- | "odd n \<Longrightarrow> even (Suc n)"
-
-code_pred (expected_modes: i => bool, o => bool) even .
-code_pred [dseq] even .
-code_pred [random_dseq] even .
-
-thm odd.equation
-thm even.equation
-thm odd.dseq_equation
-thm even.dseq_equation
-thm odd.random_dseq_equation
-thm even.random_dseq_equation
-
-values "{x. even 2}"
-values "{x. odd 2}"
-values 10 "{n. even n}"
-values 10 "{n. odd n}"
-values [expected "{}" dseq 2] "{x. even 6}"
-values [expected "{}" dseq 6] "{x. even 6}"
-values [expected "{()}" dseq 7] "{x. even 6}"
-values [dseq 2] "{x. odd 7}"
-values [dseq 6] "{x. odd 7}"
-values [dseq 7] "{x. odd 7}"
-values [expected "{()}" dseq 8] "{x. odd 7}"
-
-values [expected "{}" dseq 0] 8 "{x. even x}"
-values [expected "{0::nat}" dseq 1] 8 "{x. even x}"
-values [expected "{0::nat, 2}" dseq 3] 8 "{x. even x}"
-values [expected "{0::nat, 2}" dseq 4] 8 "{x. even x}"
-values [expected "{0::nat, 2, 4}" dseq 6] 8 "{x. even x}"
-
-values [random_dseq 1, 1, 0] 8 "{x. even x}"
-values [random_dseq 1, 1, 1] 8 "{x. even x}"
-values [random_dseq 1, 1, 2] 8 "{x. even x}"
-values [random_dseq 1, 1, 3] 8 "{x. even x}"
-values [random_dseq 1, 1, 6] 8 "{x. even x}"
-
-values [expected "{}" random_dseq 1, 1, 7] "{x. odd 7}"
-values [random_dseq 1, 1, 8] "{x. odd 7}"
-values [random_dseq 1, 1, 9] "{x. odd 7}"
-
-definition odd' where "odd' x == \<not> even x"
-
-code_pred (expected_modes: i => bool) [inductify] odd' .
-code_pred [dseq inductify] odd' .
-code_pred [random_dseq inductify] odd' .
-
-values [expected "{}" dseq 2] "{x. odd' 7}"
-values [expected "{()}" dseq 9] "{x. odd' 7}"
-values [expected "{}" dseq 2] "{x. odd' 8}"
-values [expected "{}" dseq 10] "{x. odd' 8}"
-
-
-inductive is_even :: "nat \<Rightarrow> bool"
-where
- "n mod 2 = 0 \<Longrightarrow> is_even n"
-
-code_pred (expected_modes: i => bool) is_even .
-
-subsection {* append predicate *}
-
-inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
- "append [] xs xs"
- | "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)"
-
-code_pred (modes: i => i => o => bool as "concat", o => o => i => bool as "slice", o => i => i => bool as prefix,
- i => o => i => bool as suffix, i => i => i => bool) append .
-code_pred [dseq] append .
-code_pred [random_dseq] append .
-
-thm append.equation
-thm append.dseq_equation
-thm append.random_dseq_equation
-
-values "{(ys, xs). append xs ys [0, Suc 0, 2]}"
-values "{zs. append [0, Suc 0, 2] [17, 8] zs}"
-values "{ys. append [0, Suc 0, 2] ys [0, Suc 0, 2, 17, 0, 5]}"
-
-values [expected "{}" dseq 0] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
-values [expected "{(([]::nat list), [Suc 0, 2, 3, 4, (5::nat)])}" dseq 1] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
-values [dseq 4] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
-values [dseq 6] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
-values [random_dseq 1, 1, 4] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
-values [random_dseq 1, 1, 1] 10 "{(xs, ys, zs::int list). append xs ys zs}"
-values [random_dseq 1, 1, 3] 10 "{(xs, ys, zs::int list). append xs ys zs}"
-values [random_dseq 3, 1, 3] 10 "{(xs, ys, zs::int list). append xs ys zs}"
-values [random_dseq 1, 3, 3] 10 "{(xs, ys, zs::int list). append xs ys zs}"
-values [random_dseq 1, 1, 4] 10 "{(xs, ys, zs::int list). append xs ys zs}"
-
-value [code] "Predicate.the (concat [0::int, 1, 2] [3, 4, 5])"
-value [code] "Predicate.the (slice ([]::int list))"
-
-
-text {* tricky case with alternative rules *}
-
-inductive append2
-where
- "append2 [] xs xs"
-| "append2 xs ys zs \<Longrightarrow> append2 (x # xs) ys (x # zs)"
-
-lemma append2_Nil: "append2 [] (xs::'b list) xs"
- by (simp add: append2.intros(1))
-
-lemmas [code_pred_intro] = append2_Nil append2.intros(2)
-
-code_pred (expected_modes: i => i => o => bool, o => o => i => bool, o => i => i => bool,
- i => o => i => bool, i => i => i => bool) append2
-proof -
- case append2
- from append2(1) show thesis
- proof
- fix xs
- assume "xa = []" "xb = xs" "xc = xs"
- from this append2(2) show thesis by simp
- next
- fix xs ys zs x
- assume "xa = x # xs" "xb = ys" "xc = x # zs" "append2 xs ys zs"
- from this append2(3) show thesis by fastsimp
- qed
-qed
-
-inductive tupled_append :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
-where
- "tupled_append ([], xs, xs)"
-| "tupled_append (xs, ys, zs) \<Longrightarrow> tupled_append (x # xs, ys, x # zs)"
-
-code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,
- i * o * i => bool, i * i * i => bool) tupled_append .
-code_pred [random_dseq] tupled_append .
-thm tupled_append.equation
-
-values "{xs. tupled_append ([(1::nat), 2, 3], [4, 5], xs)}"
-
-inductive tupled_append'
-where
-"tupled_append' ([], xs, xs)"
-| "[| ys = fst (xa, y); x # zs = snd (xa, y);
- tupled_append' (xs, ys, zs) |] ==> tupled_append' (x # xs, xa, y)"
-
-code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,
- i * o * i => bool, i * i * i => bool) tupled_append' .
-thm tupled_append'.equation
-
-inductive tupled_append'' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
-where
- "tupled_append'' ([], xs, xs)"
-| "ys = fst yszs ==> x # zs = snd yszs ==> tupled_append'' (xs, ys, zs) \<Longrightarrow> tupled_append'' (x # xs, yszs)"
-
-code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,
- i * o * i => bool, i * i * i => bool) tupled_append'' .
-thm tupled_append''.equation
-
-inductive tupled_append''' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
-where
- "tupled_append''' ([], xs, xs)"
-| "yszs = (ys, zs) ==> tupled_append''' (xs, yszs) \<Longrightarrow> tupled_append''' (x # xs, ys, x # zs)"
-
-code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,
- i * o * i => bool, i * i * i => bool) tupled_append''' .
-thm tupled_append'''.equation
-
-subsection {* map_ofP predicate *}
-
-inductive map_ofP :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
-where
- "map_ofP ((a, b)#xs) a b"
-| "map_ofP xs a b \<Longrightarrow> map_ofP (x#xs) a b"
-
-code_pred (expected_modes: i => o => o => bool, i => i => o => bool, i => o => i => bool, i => i => i => bool) map_ofP .
-thm map_ofP.equation
-
-subsection {* filter predicate *}
-
-inductive filter1
-for P
-where
- "filter1 P [] []"
-| "P x ==> filter1 P xs ys ==> filter1 P (x#xs) (x#ys)"
-| "\<not> P x ==> filter1 P xs ys ==> filter1 P (x#xs) ys"
-
-code_pred (expected_modes: (i => bool) => i => o => bool, (i => bool) => i => i => bool) filter1 .
-code_pred [dseq] filter1 .
-code_pred [random_dseq] filter1 .
-
-thm filter1.equation
-
-values [expected "{[0::nat, 2, 4]}"] "{xs. filter1 even [0, 1, 2, 3, 4] xs}"
-values [expected "{}" dseq 9] "{xs. filter1 even [0, 1, 2, 3, 4] xs}"
-values [expected "{[0::nat, 2, 4]}" dseq 10] "{xs. filter1 even [0, 1, 2, 3, 4] xs}"
-
-inductive filter2
-where
- "filter2 P [] []"
-| "P x ==> filter2 P xs ys ==> filter2 P (x#xs) (x#ys)"
-| "\<not> P x ==> filter2 P xs ys ==> filter2 P (x#xs) ys"
-
-code_pred (expected_modes: (i => bool) => i => i => bool, (i => bool) => i => o => bool) filter2 .
-code_pred [dseq] filter2 .
-code_pred [random_dseq] filter2 .
-
-thm filter2.equation
-thm filter2.random_dseq_equation
-
-(*
-inductive filter3
-for P
-where
- "List.filter P xs = ys ==> filter3 P xs ys"
-
-code_pred (expected_modes: (o => bool) => i => o => bool, (o => bool) => i => i => bool , (i => bool) => i => o => bool, (i => bool) => i => i => bool) [skip_proof] filter3 .
-
-code_pred [dseq] filter3 .
-thm filter3.dseq_equation
-*)
-(*
-inductive filter4
-where
- "List.filter P xs = ys ==> filter4 P xs ys"
-
-code_pred (expected_modes: i => i => o => bool, i => i => i => bool) filter4 .
-(*code_pred [depth_limited] filter4 .*)
-(*code_pred [random] filter4 .*)
-*)
-subsection {* reverse predicate *}
-
-inductive rev where
- "rev [] []"
- | "rev xs xs' ==> append xs' [x] ys ==> rev (x#xs) ys"
-
-code_pred (expected_modes: i => o => bool, o => i => bool, i => i => bool) rev .
-
-thm rev.equation
-
-values "{xs. rev [0, 1, 2, 3::nat] xs}"
-
-inductive tupled_rev where
- "tupled_rev ([], [])"
-| "tupled_rev (xs, xs') \<Longrightarrow> tupled_append (xs', [x], ys) \<Longrightarrow> tupled_rev (x#xs, ys)"
-
-code_pred (expected_modes: i * o => bool, o * i => bool, i * i => bool) tupled_rev .
-thm tupled_rev.equation
-
-subsection {* partition predicate *}
-
-inductive partition :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
- for f where
- "partition f [] [] []"
- | "f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) (x # ys) zs"
- | "\<not> f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) ys (x # zs)"
-
-code_pred (expected_modes: (i => bool) => i => o => o => bool, (i => bool) => o => i => i => bool,
- (i => bool) => i => i => o => bool, (i => bool) => i => o => i => bool, (i => bool) => i => i => i => bool)
- partition .
-code_pred [dseq] partition .
-code_pred [random_dseq] partition .
-
-values 10 "{(ys, zs). partition is_even
- [0, Suc 0, 2, 3, 4, 5, 6, 7] ys zs}"
-values 10 "{zs. partition is_even zs [0, 2] [3, 5]}"
-values 10 "{zs. partition is_even zs [0, 7] [3, 5]}"
-
-inductive tupled_partition :: "('a \<Rightarrow> bool) \<Rightarrow> ('a list \<times> 'a list \<times> 'a list) \<Rightarrow> bool"
- for f where
- "tupled_partition f ([], [], [])"
- | "f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, x # ys, zs)"
- | "\<not> f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, ys, x # zs)"
-
-code_pred (expected_modes: (i => bool) => i => bool, (i => bool) => (i * i * o) => bool, (i => bool) => (i * o * i) => bool,
- (i => bool) => (o * i * i) => bool, (i => bool) => (i * o * o) => bool) tupled_partition .
-
-thm tupled_partition.equation
-
-lemma [code_pred_intro]:
- "r a b \<Longrightarrow> tranclp r a b"
- "r a b \<Longrightarrow> tranclp r b c \<Longrightarrow> tranclp r a c"
- by auto
-
-subsection {* transitive predicate *}
-
-text {* Also look at the tabled transitive closure in the Library *}
-
-code_pred (modes: (i => o => bool) => i => i => bool, (i => o => bool) => i => o => bool as forwards_trancl,
- (o => i => bool) => i => i => bool, (o => i => bool) => o => i => bool as backwards_trancl, (o => o => bool) => i => i => bool, (o => o => bool) => i => o => bool,
- (o => o => bool) => o => i => bool, (o => o => bool) => o => o => bool) tranclp
-proof -
- case tranclp
- from this converse_tranclpE[OF this(1)] show thesis by metis
-qed
-
-
-code_pred [dseq] tranclp .
-code_pred [random_dseq] tranclp .
-thm tranclp.equation
-thm tranclp.random_dseq_equation
-
-inductive rtrancl' :: "'a => 'a => ('a => 'a => bool) => bool"
-where
- "rtrancl' x x r"
-| "r x y ==> rtrancl' y z r ==> rtrancl' x z r"
-
-code_pred [random_dseq] rtrancl' .
-
-thm rtrancl'.random_dseq_equation
-
-inductive rtrancl'' :: "('a * 'a * ('a \<Rightarrow> 'a \<Rightarrow> bool)) \<Rightarrow> bool"
-where
- "rtrancl'' (x, x, r)"
-| "r x y \<Longrightarrow> rtrancl'' (y, z, r) \<Longrightarrow> rtrancl'' (x, z, r)"
-
-code_pred rtrancl'' .
-
-inductive rtrancl''' :: "('a * ('a * 'a) * ('a * 'a => bool)) => bool"
-where
- "rtrancl''' (x, (x, x), r)"
-| "r (x, y) ==> rtrancl''' (y, (z, z), r) ==> rtrancl''' (x, (z, z), r)"
-
-code_pred rtrancl''' .
-
-
-inductive succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where
- "succ 0 1"
- | "succ m n \<Longrightarrow> succ (Suc m) (Suc n)"
-
-code_pred (modes: i => i => bool, i => o => bool, o => i => bool, o => o => bool) succ .
-code_pred [random_dseq] succ .
-thm succ.equation
-thm succ.random_dseq_equation
-
-values 10 "{(m, n). succ n m}"
-values "{m. succ 0 m}"
-values "{m. succ m 0}"
-
-text {* values command needs mode annotation of the parameter succ
-to disambiguate which mode is to be chosen. *}
-
-values [mode: i => o => bool] 20 "{n. tranclp succ 10 n}"
-values [mode: o => i => bool] 10 "{n. tranclp succ n 10}"
-values 20 "{(n, m). tranclp succ n m}"
-
-inductive example_graph :: "int => int => bool"
-where
- "example_graph 0 1"
-| "example_graph 1 2"
-| "example_graph 1 3"
-| "example_graph 4 7"
-| "example_graph 4 5"
-| "example_graph 5 6"
-| "example_graph 7 6"
-| "example_graph 7 8"
-
-inductive not_reachable_in_example_graph :: "int => int => bool"
-where "\<not> (tranclp example_graph x y) ==> not_reachable_in_example_graph x y"
-
-code_pred (expected_modes: i => i => bool) not_reachable_in_example_graph .
-
-thm not_reachable_in_example_graph.equation
-thm tranclp.equation
-value "not_reachable_in_example_graph 0 3"
-value "not_reachable_in_example_graph 4 8"
-value "not_reachable_in_example_graph 5 6"
-text {* rtrancl compilation is strange! *}
-(*
-value "not_reachable_in_example_graph 0 4"
-value "not_reachable_in_example_graph 1 6"
-value "not_reachable_in_example_graph 8 4"*)
-
-code_pred [dseq] not_reachable_in_example_graph .
-
-values [dseq 6] "{x. tranclp example_graph 0 3}"
-
-values [dseq 0] "{x. not_reachable_in_example_graph 0 3}"
-values [dseq 0] "{x. not_reachable_in_example_graph 0 4}"
-values [dseq 20] "{x. not_reachable_in_example_graph 0 4}"
-values [dseq 6] "{x. not_reachable_in_example_graph 0 3}"
-values [dseq 3] "{x. not_reachable_in_example_graph 4 2}"
-values [dseq 6] "{x. not_reachable_in_example_graph 4 2}"
-
-
-inductive not_reachable_in_example_graph' :: "int => int => bool"
-where "\<not> (rtranclp example_graph x y) ==> not_reachable_in_example_graph' x y"
-
-code_pred not_reachable_in_example_graph' .
-
-value "not_reachable_in_example_graph' 0 3"
-(* value "not_reachable_in_example_graph' 0 5" would not terminate *)
-
-
-(*values [depth_limited 0] "{x. not_reachable_in_example_graph' 0 3}"*)
-(*values [depth_limited 3] "{x. not_reachable_in_example_graph' 0 3}"*) (* fails with undefined *)
-(*values [depth_limited 5] "{x. not_reachable_in_example_graph' 0 3}"*)
-(*values [depth_limited 1] "{x. not_reachable_in_example_graph' 0 4}"*)
-(*values [depth_limit = 4] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
-(*values [depth_limit = 20] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
-
-code_pred [dseq] not_reachable_in_example_graph' .
-
-(*thm not_reachable_in_example_graph'.dseq_equation*)
-
-(*values [dseq 0] "{x. not_reachable_in_example_graph' 0 3}"*)
-(*values [depth_limited 3] "{x. not_reachable_in_example_graph' 0 3}"*) (* fails with undefined *)
-(*values [depth_limited 5] "{x. not_reachable_in_example_graph' 0 3}"
-values [depth_limited 1] "{x. not_reachable_in_example_graph' 0 4}"*)
-(*values [depth_limit = 4] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
-(*values [depth_limit = 20] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
-
-
-subsection {* IMP *}
-
-types
- var = nat
- state = "int list"
-
-datatype com =
- Skip |
- Ass var "state => int" |
- Seq com com |
- IF "state => bool" com com |
- While "state => bool" com
-
-inductive exec :: "com => state => state => bool" where
-"exec Skip s s" |
-"exec (Ass x e) s (s[x := e(s)])" |
-"exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" |
-"b s ==> exec c1 s t ==> exec (IF b c1 c2) s t" |
-"~b s ==> exec c2 s t ==> exec (IF b c1 c2) s t" |
-"~b s ==> exec (While b c) s s" |
-"b s1 ==> exec c s1 s2 ==> exec (While b c) s2 s3 ==> exec (While b c) s1 s3"
-
-code_pred exec .
-
-values "{t. exec
- (While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1))))
- [3,5] t}"
-
-
-inductive tupled_exec :: "(com \<times> state \<times> state) \<Rightarrow> bool" where
-"tupled_exec (Skip, s, s)" |
-"tupled_exec (Ass x e, s, s[x := e(s)])" |
-"tupled_exec (c1, s1, s2) ==> tupled_exec (c2, s2, s3) ==> tupled_exec (Seq c1 c2, s1, s3)" |
-"b s ==> tupled_exec (c1, s, t) ==> tupled_exec (IF b c1 c2, s, t)" |
-"~b s ==> tupled_exec (c2, s, t) ==> tupled_exec (IF b c1 c2, s, t)" |
-"~b s ==> tupled_exec (While b c, s, s)" |
-"b s1 ==> tupled_exec (c, s1, s2) ==> tupled_exec (While b c, s2, s3) ==> tupled_exec (While b c, s1, s3)"
-
-code_pred tupled_exec .
-
-values "{s. tupled_exec (While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1))), [3, 5], s)}"
-
-subsection {* CCS *}
-
-text{* This example formalizes finite CCS processes without communication or
-recursion. For simplicity, labels are natural numbers. *}
-
-datatype proc = nil | pre nat proc | or proc proc | par proc proc
-
-inductive step :: "proc \<Rightarrow> nat \<Rightarrow> proc \<Rightarrow> bool" where
-"step (pre n p) n p" |
-"step p1 a q \<Longrightarrow> step (or p1 p2) a q" |
-"step p2 a q \<Longrightarrow> step (or p1 p2) a q" |
-"step p1 a q \<Longrightarrow> step (par p1 p2) a (par q p2)" |
-"step p2 a q \<Longrightarrow> step (par p1 p2) a (par p1 q)"
-
-code_pred step .
-
-inductive steps where
-"steps p [] p" |
-"step p a q \<Longrightarrow> steps q as r \<Longrightarrow> steps p (a#as) r"
-
-code_pred steps .
-
-values 3
- "{as . steps (par (or (pre 0 nil) (pre 1 nil)) (pre 2 nil)) as (par nil nil)}"
-
-values 5
- "{as . steps (par (or (pre 0 nil) (pre 1 nil)) (pre 2 nil)) as (par nil nil)}"
-
-values 3 "{(a,q). step (par nil nil) a q}"
-
-
-inductive tupled_step :: "(proc \<times> nat \<times> proc) \<Rightarrow> bool"
-where
-"tupled_step (pre n p, n, p)" |
-"tupled_step (p1, a, q) \<Longrightarrow> tupled_step (or p1 p2, a, q)" |
-"tupled_step (p2, a, q) \<Longrightarrow> tupled_step (or p1 p2, a, q)" |
-"tupled_step (p1, a, q) \<Longrightarrow> tupled_step (par p1 p2, a, par q p2)" |
-"tupled_step (p2, a, q) \<Longrightarrow> tupled_step (par p1 p2, a, par p1 q)"
-
-code_pred tupled_step .
-thm tupled_step.equation
-
-subsection {* divmod *}
-
-inductive divmod_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
- "k < l \<Longrightarrow> divmod_rel k l 0 k"
- | "k \<ge> l \<Longrightarrow> divmod_rel (k - l) l q r \<Longrightarrow> divmod_rel k l (Suc q) r"
-
-code_pred divmod_rel ..
-thm divmod_rel.equation
-value [code] "Predicate.the (divmod_rel_i_i_o_o 1705 42)"
-
-subsection {* Transforming predicate logic into logic programs *}
-
-subsection {* Transforming functions into logic programs *}
-definition
- "case_f xs ys = (case (xs @ ys) of [] => [] | (x # xs) => xs)"
-
-code_pred [inductify] case_f .
-thm case_fP.equation
-thm case_fP.intros
-
-fun fold_map_idx where
- "fold_map_idx f i y [] = (y, [])"
-| "fold_map_idx f i y (x # xs) =
- (let (y', x') = f i y x; (y'', xs') = fold_map_idx f (Suc i) y' xs
- in (y'', x' # xs'))"
-
-text {* mode analysis explores thousand modes - this is infeasible at the moment... *}
-(*code_pred [inductify, show_steps] fold_map_idx .*)
-
-subsection {* Minimum *}
-
-definition Min
-where "Min s r x \<equiv> s x \<and> (\<forall>y. r x y \<longrightarrow> x = y)"
-
-code_pred [inductify] Min .
-thm Min.equation
-
-subsection {* Lexicographic order *}
-
-declare lexord_def[code_pred_def]
-code_pred [inductify] lexord .
-code_pred [random_dseq inductify] lexord .
-
-thm lexord.equation
-thm lexord.random_dseq_equation
-
-inductive less_than_nat :: "nat * nat => bool"
-where
- "less_than_nat (0, x)"
-| "less_than_nat (x, y) ==> less_than_nat (Suc x, Suc y)"
-
-code_pred less_than_nat .
-
-code_pred [dseq] less_than_nat .
-code_pred [random_dseq] less_than_nat .
-
-inductive test_lexord :: "nat list * nat list => bool"
-where
- "lexord less_than_nat (xs, ys) ==> test_lexord (xs, ys)"
-
-code_pred test_lexord .
-code_pred [dseq] test_lexord .
-code_pred [random_dseq] test_lexord .
-thm test_lexord.dseq_equation
-thm test_lexord.random_dseq_equation
-
-values "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"
-(*values [depth_limited 5] "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"*)
-
-declare list.size(3,4)[code_pred_def]
-lemmas [code_pred_def] = lexn_conv lex_conv lenlex_conv
-(*
-code_pred [inductify] lexn .
-thm lexn.equation
-*)
-(*
-code_pred [random_dseq inductify] lexn .
-thm lexn.random_dseq_equation
-
-values [random_dseq 4, 4, 6] 100 "{(n, xs, ys::int list). lexn (%(x, y). x <= y) n (xs, ys)}"
-*)
-inductive has_length
-where
- "has_length [] 0"
-| "has_length xs i ==> has_length (x # xs) (Suc i)"
-
-lemma has_length:
- "has_length xs n = (length xs = n)"
-proof (rule iffI)
- assume "has_length xs n"
- from this show "length xs = n"
- by (rule has_length.induct) auto
-next
- assume "length xs = n"
- from this show "has_length xs n"
- by (induct xs arbitrary: n) (auto intro: has_length.intros)
-qed
-
-lemma lexn_intros [code_pred_intro]:
- "has_length xs i ==> has_length ys i ==> r (x, y) ==> lexn r (Suc i) (x # xs, y # ys)"
- "lexn r i (xs, ys) ==> lexn r (Suc i) (x # xs, x # ys)"
-proof -
- assume "has_length xs i" "has_length ys i" "r (x, y)"
- from this has_length show "lexn r (Suc i) (x # xs, y # ys)"
- unfolding lexn_conv Collect_def mem_def
- by fastsimp
-next
- assume "lexn r i (xs, ys)"
- thm lexn_conv
- from this show "lexn r (Suc i) (x#xs, x#ys)"
- unfolding Collect_def mem_def lexn_conv
- apply auto
- apply (rule_tac x="x # xys" in exI)
- by auto
-qed
-
-code_pred [random_dseq inductify] lexn
-proof -
- fix r n xs ys
- assume 1: "lexn r n (xs, ys)"
- assume 2: "\<And>r' i x xs' y ys'. r = r' ==> n = Suc i ==> (xs, ys) = (x # xs', y # ys') ==> has_length xs' i ==> has_length ys' i ==> r' (x, y) ==> thesis"
- assume 3: "\<And>r' i x xs' ys'. r = r' ==> n = Suc i ==> (xs, ys) = (x # xs', x # ys') ==> lexn r' i (xs', ys') ==> thesis"
- from 1 2 3 show thesis
- unfolding lexn_conv Collect_def mem_def
- apply (auto simp add: has_length)
- apply (case_tac xys)
- apply auto
- apply fastsimp
- apply fastsimp done
-qed
-
-
-values [random_dseq 1, 2, 5] 10 "{(n, xs, ys::int list). lexn (%(x, y). x <= y) n (xs, ys)}"
-thm lenlex_conv
-thm lex_conv
-declare list.size(3,4)[code_pred_def]
-(*code_pred [inductify, show_steps, show_intermediate_results] length .*)
-setup {* Predicate_Compile_Data.ignore_consts [@{const_name Orderings.top_class.top}] *}
-code_pred [inductify] lex .
-thm lex.equation
-thm lex_def
-declare lenlex_conv[code_pred_def]
-code_pred [inductify] lenlex .
-thm lenlex.equation
-
-code_pred [random_dseq inductify] lenlex .
-thm lenlex.random_dseq_equation
-
-values [random_dseq 4, 2, 4] 100 "{(xs, ys::int list). lenlex (%(x, y). x <= y) (xs, ys)}"
-thm lists.intros
-
-code_pred [inductify] lists .
-thm lists.equation
-
-subsection {* AVL Tree *}
-
-datatype 'a tree = ET | MKT 'a "'a tree" "'a tree" nat
-fun height :: "'a tree => nat" where
-"height ET = 0"
-| "height (MKT x l r h) = max (height l) (height r) + 1"
-
-consts avl :: "'a tree => bool"
-primrec
- "avl ET = True"
- "avl (MKT x l r h) = ((height l = height r \<or> height l = 1 + height r \<or> height r = 1+height l) \<and>
- h = max (height l) (height r) + 1 \<and> avl l \<and> avl r)"
-(*
-code_pred [inductify] avl .
-thm avl.equation*)
-
-code_pred [random_dseq inductify] avl .
-thm avl.random_dseq_equation
-
-values [random_dseq 2, 1, 7] 5 "{t:: int tree. avl t}"
-
-fun set_of
-where
-"set_of ET = {}"
-| "set_of (MKT n l r h) = insert n (set_of l \<union> set_of r)"
-
-fun is_ord :: "nat tree => bool"
-where
-"is_ord ET = True"
-| "is_ord (MKT n l r h) =
- ((\<forall>n' \<in> set_of l. n' < n) \<and> (\<forall>n' \<in> set_of r. n < n') \<and> is_ord l \<and> is_ord r)"
-
-code_pred (expected_modes: i => o => bool, i => i => bool) [inductify] set_of .
-thm set_of.equation
-
-code_pred (expected_modes: i => bool) [inductify] is_ord .
-thm is_ord_aux.equation
-thm is_ord.equation
-
-
-subsection {* Definitions about Relations *}
-term "converse"
-code_pred (modes:
- (i * i => bool) => i * i => bool,
- (i * o => bool) => o * i => bool,
- (i * o => bool) => i * i => bool,
- (o * i => bool) => i * o => bool,
- (o * i => bool) => i * i => bool,
- (o * o => bool) => o * o => bool,
- (o * o => bool) => i * o => bool,
- (o * o => bool) => o * i => bool,
- (o * o => bool) => i * i => bool) [inductify] converse .
-
-thm converse.equation
-code_pred [inductify] rel_comp .
-thm rel_comp.equation
-code_pred [inductify] Image .
-thm Image.equation
-declare singleton_iff[code_pred_inline]
-declare Id_on_def[unfolded Bex_def UNION_def singleton_iff, code_pred_def]
-
-code_pred (expected_modes:
- (o => bool) => o => bool,
- (o => bool) => i * o => bool,
- (o => bool) => o * i => bool,
- (o => bool) => i => bool,
- (i => bool) => i * o => bool,
- (i => bool) => o * i => bool,
- (i => bool) => i => bool) [inductify] Id_on .
-thm Id_on.equation
-thm Domain_def
-code_pred (modes:
- (o * o => bool) => o => bool,
- (o * o => bool) => i => bool,
- (i * o => bool) => i => bool) [inductify] Domain .
-thm Domain.equation
-
-thm Range_def
-code_pred (modes:
- (o * o => bool) => o => bool,
- (o * o => bool) => i => bool,
- (o * i => bool) => i => bool) [inductify] Range .
-thm Range.equation
-
-code_pred [inductify] Field .
-thm Field.equation
-
-thm refl_on_def
-code_pred [inductify] refl_on .
-thm refl_on.equation
-code_pred [inductify] total_on .
-thm total_on.equation
-code_pred [inductify] antisym .
-thm antisym.equation
-code_pred [inductify] trans .
-thm trans.equation
-code_pred [inductify] single_valued .
-thm single_valued.equation
-thm inv_image_def
-code_pred [inductify] inv_image .
-thm inv_image.equation
-
-subsection {* Inverting list functions *}
-
-(*code_pred [inductify] length .
-code_pred [random inductify] length .
-thm size_listP.equation
-thm size_listP.random_equation
-*)
-(*values [random] 1 "{xs. size_listP (xs::nat list) (5::nat)}"*)
-
-code_pred (expected_modes: i => o => bool, o => i => bool, i => i => bool) [inductify] List.concat .
-thm concatP.equation
-
-values "{ys. concatP [[1, 2], [3, (4::int)]] ys}"
-values "{ys. concatP [[1, 2], [3]] [1, 2, (3::nat)]}"
-
-code_pred [dseq inductify] List.concat .
-thm concatP.dseq_equation
-
-values [dseq 3] 3
- "{xs. concatP xs ([0] :: nat list)}"
-
-values [dseq 5] 3
- "{xs. concatP xs ([1] :: int list)}"
-
-values [dseq 5] 3
- "{xs. concatP xs ([1] :: nat list)}"
-
-values [dseq 5] 3
- "{xs. concatP xs [(1::int), 2]}"
-
-code_pred (expected_modes: i => o => bool, i => i => bool) [inductify] hd .
-thm hdP.equation
-values "{x. hdP [1, 2, (3::int)] x}"
-values "{(xs, x). hdP [1, 2, (3::int)] 1}"
-
-code_pred (expected_modes: i => o => bool, i => i => bool) [inductify] tl .
-thm tlP.equation
-values "{x. tlP [1, 2, (3::nat)] x}"
-values "{x. tlP [1, 2, (3::int)] [3]}"
-
-code_pred [inductify] last .
-thm lastP.equation
-
-code_pred [inductify] butlast .
-thm butlastP.equation
-
-code_pred [inductify] take .
-thm takeP.equation
-
-code_pred [inductify] drop .
-thm dropP.equation
-code_pred [inductify] zip .
-thm zipP.equation
-
-code_pred [inductify] upt .
-code_pred [inductify] remdups .
-thm remdupsP.equation
-code_pred [dseq inductify] remdups .
-values [dseq 4] 5 "{xs. remdupsP xs [1, (2::int)]}"
-
-code_pred [inductify] remove1 .
-thm remove1P.equation
-values "{xs. remove1P 1 xs [2, (3::int)]}"
-
-code_pred [inductify] removeAll .
-thm removeAllP.equation
-code_pred [dseq inductify] removeAll .
-
-values [dseq 4] 10 "{xs. removeAllP 1 xs [(2::nat)]}"
-
-code_pred [inductify] distinct .
-thm distinct.equation
-code_pred [inductify] replicate .
-thm replicateP.equation
-values 5 "{(n, xs). replicateP n (0::int) xs}"
-
-code_pred [inductify] splice .
-thm splice.simps
-thm spliceP.equation
-
-values "{xs. spliceP xs [1, 2, 3] [1, 1, 1, 2, 1, (3::nat)]}"
-
-code_pred [inductify] List.rev .
-code_pred [inductify] map .
-code_pred [inductify] foldr .
-code_pred [inductify] foldl .
-code_pred [inductify] filter .
-code_pred [random_dseq inductify] filter .
-
-subsection {* Context Free Grammar *}
-
-datatype alphabet = a | b
-
-inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
- "[] \<in> S\<^isub>1"
-| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
-| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
-| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
-| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
-| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
-
-code_pred [inductify] S\<^isub>1p .
-code_pred [random_dseq inductify] S\<^isub>1p .
-thm S\<^isub>1p.equation
-thm S\<^isub>1p.random_dseq_equation
-
-values [random_dseq 5, 5, 5] 5 "{x. S\<^isub>1p x}"
-
-inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
- "[] \<in> S\<^isub>2"
-| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
-| "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2"
-| "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
-| "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
-| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
-
-code_pred [random_dseq inductify] S\<^isub>2p .
-thm S\<^isub>2p.random_dseq_equation
-thm A\<^isub>2p.random_dseq_equation
-thm B\<^isub>2p.random_dseq_equation
-
-values [random_dseq 5, 5, 5] 10 "{x. S\<^isub>2p x}"
-
-inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
- "[] \<in> S\<^isub>3"
-| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3"
-| "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3"
-| "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3"
-| "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3"
-| "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3"
-
-code_pred [inductify] S\<^isub>3p .
-thm S\<^isub>3p.equation
-
-values 10 "{x. S\<^isub>3p x}"
-
-inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
- "[] \<in> S\<^isub>4"
-| "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4"
-| "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4"
-| "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4"
-| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
-| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
-| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
-
-code_pred (expected_modes: o => bool, i => bool) S\<^isub>4p .
-
-subsection {* Lambda *}
-
-datatype type =
- Atom nat
- | Fun type type (infixr "\<Rightarrow>" 200)
-
-datatype dB =
- Var nat
- | App dB dB (infixl "\<degree>" 200)
- | Abs type dB
-
-primrec
- nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
-where
- "[]\<langle>i\<rangle> = None"
-| "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
-
-inductive nth_el' :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool"
-where
- "nth_el' (x # xs) 0 x"
-| "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y"
-
-inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" ("_ \<turnstile> _ : _" [50, 50, 50] 50)
- where
- Var [intro!]: "nth_el' env x T \<Longrightarrow> env \<turnstile> Var x : T"
- | Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)"
- | App [intro!]: "env \<turnstile> s : T \<Rightarrow> U \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U"
-
-primrec
- lift :: "[dB, nat] => dB"
-where
- "lift (Var i) k = (if i < k then Var i else Var (i + 1))"
- | "lift (s \<degree> t) k = lift s k \<degree> lift t k"
- | "lift (Abs T s) k = Abs T (lift s (k + 1))"
-
-primrec
- subst :: "[dB, dB, nat] => dB" ("_[_'/_]" [300, 0, 0] 300)
-where
- subst_Var: "(Var i)[s/k] =
- (if k < i then Var (i - 1) else if i = k then s else Var i)"
- | subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
- | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"
-
-inductive beta :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<beta>" 50)
- where
- beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
- | appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
- | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
- | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"
-
-code_pred (expected_modes: i => i => o => bool, i => i => i => bool) typing .
-thm typing.equation
-
-code_pred (modes: i => i => bool, i => o => bool as reduce') beta .
-thm beta.equation
-
-values "{x. App (Abs (Atom 0) (Var 0)) (Var 1) \<rightarrow>\<^sub>\<beta> x}"
-
-definition "reduce t = Predicate.the (reduce' t)"
-
-value "reduce (App (Abs (Atom 0) (Var 0)) (Var 1))"
-
-code_pred [dseq] typing .
-code_pred [random_dseq] typing .
-
-values [random_dseq 1,1,5] 10 "{(\<Gamma>, t, T). \<Gamma> \<turnstile> t : T}"
-
-subsection {* A minimal example of yet another semantics *}
-
-text {* thanks to Elke Salecker *}
-
-types
- vname = nat
- vvalue = int
- var_assign = "vname \<Rightarrow> vvalue" --"variable assignment"
-
-datatype ir_expr =
- IrConst vvalue
-| ObjAddr vname
-| Add ir_expr ir_expr
-
-datatype val =
- IntVal vvalue
-
-record configuration =
- Env :: var_assign
-
-inductive eval_var ::
- "ir_expr \<Rightarrow> configuration \<Rightarrow> val \<Rightarrow> bool"
-where
- irconst: "eval_var (IrConst i) conf (IntVal i)"
-| objaddr: "\<lbrakk> Env conf n = i \<rbrakk> \<Longrightarrow> eval_var (ObjAddr n) conf (IntVal i)"
-| plus: "\<lbrakk> eval_var l conf (IntVal vl); eval_var r conf (IntVal vr) \<rbrakk> \<Longrightarrow> eval_var (Add l r) conf (IntVal (vl+vr))"
-
-
-code_pred eval_var .
-thm eval_var.equation
-
-values "{val. eval_var (Add (IrConst 1) (IrConst 2)) (| Env = (\<lambda>x. 0)|) val}"
-
-end
--- a/src/HOL/ex/ROOT.ML Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOL/ex/ROOT.ML Thu Mar 25 17:56:31 2010 +0100
@@ -11,10 +11,7 @@
"Eval_Examples",
"Codegenerator_Test",
"Codegenerator_Pretty_Test",
- "NormalForm",
- "Predicate_Compile",
- "Predicate_Compile_ex",
- "Predicate_Compile_Quickcheck"
+ "NormalForm"
];
use_thys [
--- a/src/HOLCF/Algebraic.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Algebraic.thy Thu Mar 25 17:56:31 2010 +0100
@@ -714,7 +714,8 @@
apply (rule finite_deflation_approx)
apply (rule chainE)
apply (rule chain_approx)
-apply (simp add: cast_alg_defl_principal Abs_fin_defl_inverse finite_deflation_approx)
+apply (simp add: cast_alg_defl_principal
+ Abs_fin_defl_inverse finite_deflation_approx)
done
subsection {* Bifinite domains and algebraic deflations *}
--- a/src/HOLCF/Cfun.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Cfun.thy Thu Mar 25 17:56:31 2010 +0100
@@ -185,23 +185,20 @@
done
lemmas monofun_Rep_CFun = cont_Rep_CFun [THEN cont2mono]
-lemmas contlub_Rep_CFun = cont_Rep_CFun [THEN cont2contlub]
lemmas monofun_Rep_CFun1 = cont_Rep_CFun1 [THEN cont2mono, standard]
-lemmas contlub_Rep_CFun1 = cont_Rep_CFun1 [THEN cont2contlub, standard]
lemmas monofun_Rep_CFun2 = cont_Rep_CFun2 [THEN cont2mono, standard]
-lemmas contlub_Rep_CFun2 = cont_Rep_CFun2 [THEN cont2contlub, standard]
text {* contlub, cont properties of @{term Rep_CFun} in each argument *}
lemma contlub_cfun_arg: "chain Y \<Longrightarrow> f\<cdot>(\<Squnion>i. Y i) = (\<Squnion>i. f\<cdot>(Y i))"
-by (rule contlub_Rep_CFun2 [THEN contlubE])
+by (rule cont_Rep_CFun2 [THEN cont2contlubE])
lemma cont_cfun_arg: "chain Y \<Longrightarrow> range (\<lambda>i. f\<cdot>(Y i)) <<| f\<cdot>(\<Squnion>i. Y i)"
by (rule cont_Rep_CFun2 [THEN contE])
lemma contlub_cfun_fun: "chain F \<Longrightarrow> (\<Squnion>i. F i)\<cdot>x = (\<Squnion>i. F i\<cdot>x)"
-by (rule contlub_Rep_CFun1 [THEN contlubE])
+by (rule cont_Rep_CFun1 [THEN cont2contlubE])
lemma cont_cfun_fun: "chain F \<Longrightarrow> range (\<lambda>i. F i\<cdot>x) <<| (\<Squnion>i. F i)\<cdot>x"
by (rule cont_Rep_CFun1 [THEN contE])
@@ -530,26 +527,16 @@
apply (auto simp add: monofun_cfun_arg)
done
-(*FIXME: long proof*)
-lemma contlub_strictify2: "contlub (\<lambda>x. if x = \<bottom> then \<bottom> else f\<cdot>x)"
-apply (rule contlubI)
-apply (case_tac "(\<Squnion>i. Y i) = \<bottom>")
-apply (drule (1) chain_UU_I)
-apply simp
-apply (simp del: if_image_distrib)
-apply (simp only: contlub_cfun_arg)
-apply (rule lub_equal2)
-apply (rule chain_mono2 [THEN exE])
-apply (erule chain_UU_I_inverse2)
-apply (assumption)
-apply (rule_tac x=x in exI, clarsimp)
-apply (erule chain_monofun)
-apply (erule monofun_strictify2 [THEN ch2ch_monofun])
+lemma cont_strictify2: "cont (\<lambda>x. if x = \<bottom> then \<bottom> else f\<cdot>x)"
+apply (rule contI2)
+apply (rule monofun_strictify2)
+apply (case_tac "(\<Squnion>i. Y i) = \<bottom>", simp)
+apply (simp add: contlub_cfun_arg del: if_image_distrib)
+apply (drule chain_UU_I_inverse2, clarify, rename_tac j)
+apply (rule lub_mono2, rule_tac x=j in exI, simp_all)
+apply (auto dest!: chain_mono_less)
done
-lemmas cont_strictify2 =
- monocontlub2cont [OF monofun_strictify2 contlub_strictify2, standard]
-
lemma strictify_conv_if: "strictify\<cdot>f\<cdot>x = (if x = \<bottom> then \<bottom> else f\<cdot>x)"
unfolding strictify_def
by (simp add: cont_strictify1 cont_strictify2 cont2cont_LAM)
@@ -560,17 +547,26 @@
lemma strictify2 [simp]: "x \<noteq> \<bottom> \<Longrightarrow> strictify\<cdot>f\<cdot>x = f\<cdot>x"
by (simp add: strictify_conv_if)
-subsection {* Continuous let-bindings *}
+subsection {* Continuity of let-bindings *}
-definition
- CLet :: "'a \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'b" where
- "CLet = (\<Lambda> s f. f\<cdot>s)"
+lemma cont2cont_Let:
+ assumes f: "cont (\<lambda>x. f x)"
+ assumes g1: "\<And>y. cont (\<lambda>x. g x y)"
+ assumes g2: "\<And>x. cont (\<lambda>y. g x y)"
+ shows "cont (\<lambda>x. let y = f x in g x y)"
+unfolding Let_def using f g2 g1 by (rule cont_apply)
-syntax
- "_CLet" :: "[letbinds, 'a] => 'a" ("(Let (_)/ in (_))" 10)
-
-translations
- "_CLet (_binds b bs) e" == "_CLet b (_CLet bs e)"
- "Let x = a in e" == "CONST CLet\<cdot>a\<cdot>(\<Lambda> x. e)"
+lemma cont2cont_Let' [cont2cont]:
+ assumes f: "cont (\<lambda>x. f x)"
+ assumes g: "cont (\<lambda>p. g (fst p) (snd p))"
+ shows "cont (\<lambda>x. let y = f x in g x y)"
+using f
+proof (rule cont2cont_Let)
+ fix x show "cont (\<lambda>y. g x y)"
+ using g by (rule cont_fst_snd_D2)
+next
+ fix y show "cont (\<lambda>x. g x y)"
+ using g by (rule cont_fst_snd_D1)
+qed
end
--- a/src/HOLCF/CompactBasis.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/CompactBasis.thy Thu Mar 25 17:56:31 2010 +0100
@@ -217,7 +217,8 @@
apply (cut_tac a=x in PDUnit)
apply (simp add: PDUnit_def)
apply (drule_tac a=x in PDPlus)
-apply (simp add: PDUnit_def PDPlus_def Abs_pd_basis_inverse [unfolded pd_basis_def])
+apply (simp add: PDUnit_def PDPlus_def
+ Abs_pd_basis_inverse [unfolded pd_basis_def])
done
lemma pd_basis_induct:
@@ -245,7 +246,8 @@
shows "fold_pd g f (PDPlus t u) = f (fold_pd g f t) (fold_pd g f u)"
proof -
interpret ab_semigroup_idem_mult f by fact
- show ?thesis unfolding fold_pd_def Rep_PDPlus by (simp add: image_Un fold1_Un2)
+ show ?thesis unfolding fold_pd_def Rep_PDPlus
+ by (simp add: image_Un fold1_Un2)
qed
text {* Take function for powerdomain basis *}
--- a/src/HOLCF/Cont.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Cont.thy Thu Mar 25 17:56:31 2010 +0100
@@ -22,22 +22,17 @@
monofun :: "('a \<Rightarrow> 'b) \<Rightarrow> bool" -- "monotonicity" where
"monofun f = (\<forall>x y. x \<sqsubseteq> y \<longrightarrow> f x \<sqsubseteq> f y)"
+(*
definition
contlub :: "('a::cpo \<Rightarrow> 'b::cpo) \<Rightarrow> bool" -- "first cont. def" where
"contlub f = (\<forall>Y. chain Y \<longrightarrow> f (\<Squnion>i. Y i) = (\<Squnion>i. f (Y i)))"
+*)
definition
- cont :: "('a::cpo \<Rightarrow> 'b::cpo) \<Rightarrow> bool" -- "secnd cont. def" where
+ cont :: "('a::cpo \<Rightarrow> 'b::cpo) \<Rightarrow> bool"
+where
"cont f = (\<forall>Y. chain Y \<longrightarrow> range (\<lambda>i. f (Y i)) <<| f (\<Squnion>i. Y i))"
-lemma contlubI:
- "\<lbrakk>\<And>Y. chain Y \<Longrightarrow> f (\<Squnion>i. Y i) = (\<Squnion>i. f (Y i))\<rbrakk> \<Longrightarrow> contlub f"
-by (simp add: contlub_def)
-
-lemma contlubE:
- "\<lbrakk>contlub f; chain Y\<rbrakk> \<Longrightarrow> f (\<Squnion>i. Y i) = (\<Squnion>i. f (Y i))"
-by (simp add: contlub_def)
-
lemma contI:
"\<lbrakk>\<And>Y. chain Y \<Longrightarrow> range (\<lambda>i. f (Y i)) <<| f (\<Squnion>i. Y i)\<rbrakk> \<Longrightarrow> cont f"
by (simp add: cont_def)
@@ -55,7 +50,7 @@
by (simp add: monofun_def)
-subsection {* @{prop "monofun f \<and> contlub f \<equiv> cont f"} *}
+subsection {* Equivalence of alternate definition *}
text {* monotone functions map chains to chains *}
@@ -74,16 +69,7 @@
apply (erule ub_rangeD)
done
-text {* left to right: @{prop "monofun f \<and> contlub f \<Longrightarrow> cont f"} *}
-
-lemma monocontlub2cont: "\<lbrakk>monofun f; contlub f\<rbrakk> \<Longrightarrow> cont f"
-apply (rule contI)
-apply (rule thelubE)
-apply (erule (1) ch2ch_monofun)
-apply (erule (1) contlubE [symmetric])
-done
-
-text {* first a lemma about binary chains *}
+text {* a lemma about binary chains *}
lemma binchain_cont:
"\<lbrakk>cont f; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> range (\<lambda>i::nat. f (if i = 0 then x else y)) <<| f y"
@@ -95,8 +81,7 @@
apply (erule lub_bin_chain [THEN thelubI])
done
-text {* right to left: @{prop "cont f \<Longrightarrow> monofun f \<and> contlub f"} *}
-text {* part1: @{prop "cont f \<Longrightarrow> monofun f"} *}
+text {* continuity implies monotonicity *}
lemma cont2mono: "cont f \<Longrightarrow> monofun f"
apply (rule monofunI)
@@ -109,33 +94,30 @@
lemmas ch2ch_cont = cont2mono [THEN ch2ch_monofun]
-text {* right to left: @{prop "cont f \<Longrightarrow> monofun f \<and> contlub f"} *}
-text {* part2: @{prop "cont f \<Longrightarrow> contlub f"} *}
+text {* continuity implies preservation of lubs *}
-lemma cont2contlub: "cont f \<Longrightarrow> contlub f"
-apply (rule contlubI)
+lemma cont2contlubE:
+ "\<lbrakk>cont f; chain Y\<rbrakk> \<Longrightarrow> f (\<Squnion> i. Y i) = (\<Squnion> i. f (Y i))"
apply (rule thelubI [symmetric])
apply (erule (1) contE)
done
-lemmas cont2contlubE = cont2contlub [THEN contlubE]
-
lemma contI2:
assumes mono: "monofun f"
assumes below: "\<And>Y. \<lbrakk>chain Y; chain (\<lambda>i. f (Y i))\<rbrakk>
\<Longrightarrow> f (\<Squnion>i. Y i) \<sqsubseteq> (\<Squnion>i. f (Y i))"
shows "cont f"
-apply (rule monocontlub2cont)
-apply (rule mono)
-apply (rule contlubI)
+apply (rule contI)
+apply (rule thelubE)
+apply (erule ch2ch_monofun [OF mono])
apply (rule below_antisym)
-apply (rule below, assumption)
-apply (erule ch2ch_monofun [OF mono])
apply (rule is_lub_thelub)
apply (erule ch2ch_monofun [OF mono])
apply (rule ub2ub_monofun [OF mono])
apply (rule is_lubD1)
apply (erule cpo_lubI)
+apply (rule below, assumption)
+apply (erule ch2ch_monofun [OF mono])
done
subsection {* Continuity simproc *}
@@ -190,7 +172,7 @@
assumes 2: "\<And>x. cont (\<lambda>y. f x y)"
assumes 3: "\<And>y. cont (\<lambda>x. f x y)"
shows "cont (\<lambda>x. (f x) (t x))"
-proof (rule monocontlub2cont [OF monofunI contlubI])
+proof (rule contI2 [OF monofunI])
fix x y :: "'a" assume "x \<sqsubseteq> y"
then show "f x (t x) \<sqsubseteq> f y (t y)"
by (auto intro: cont2monofunE [OF 1]
@@ -199,11 +181,11 @@
below_trans)
next
fix Y :: "nat \<Rightarrow> 'a" assume "chain Y"
- then show "f (\<Squnion>i. Y i) (t (\<Squnion>i. Y i)) = (\<Squnion>i. f (Y i) (t (Y i)))"
+ then show "f (\<Squnion>i. Y i) (t (\<Squnion>i. Y i)) \<sqsubseteq> (\<Squnion>i. f (Y i) (t (Y i)))"
by (simp only: cont2contlubE [OF 1] ch2ch_cont [OF 1]
cont2contlubE [OF 2] ch2ch_cont [OF 2]
cont2contlubE [OF 3] ch2ch_cont [OF 3]
- diag_lub)
+ diag_lub below_refl)
qed
lemma cont_compose:
@@ -234,9 +216,7 @@
by (rule cont2mono [THEN monofun_finch2finch])
lemma chfindom_monofun2cont: "monofun f \<Longrightarrow> cont (f::'a::chfin \<Rightarrow> 'b::cpo)"
-apply (rule monocontlub2cont)
-apply assumption
-apply (rule contlubI)
+apply (erule contI2)
apply (frule chfin2finch)
apply (clarsimp simp add: finite_chain_def)
apply (subgoal_tac "max_in_chain i (\<lambda>i. f (Y i))")
--- a/src/HOLCF/ConvexPD.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/ConvexPD.thy Thu Mar 25 17:56:31 2010 +0100
@@ -517,16 +517,19 @@
lemma convex_map_approx: "convex_map\<cdot>(approx n)\<cdot>xs = approx n\<cdot>xs"
by (induct xs rule: convex_pd_induct, simp_all)
-lemma ep_pair_convex_map: "ep_pair e p \<Longrightarrow> ep_pair (convex_map\<cdot>e) (convex_map\<cdot>p)"
+lemma ep_pair_convex_map:
+ "ep_pair e p \<Longrightarrow> ep_pair (convex_map\<cdot>e) (convex_map\<cdot>p)"
apply default
apply (induct_tac x rule: convex_pd_induct, simp_all add: ep_pair.e_inverse)
-apply (induct_tac y rule: convex_pd_induct, simp_all add: ep_pair.e_p_below monofun_cfun)
+apply (induct_tac y rule: convex_pd_induct)
+apply (simp_all add: ep_pair.e_p_below monofun_cfun)
done
lemma deflation_convex_map: "deflation d \<Longrightarrow> deflation (convex_map\<cdot>d)"
apply default
apply (induct_tac x rule: convex_pd_induct, simp_all add: deflation.idem)
-apply (induct_tac x rule: convex_pd_induct, simp_all add: deflation.below monofun_cfun)
+apply (induct_tac x rule: convex_pd_induct)
+apply (simp_all add: deflation.below monofun_cfun)
done
--- a/src/HOLCF/Cprod.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Cprod.thy Thu Mar 25 17:56:31 2010 +0100
@@ -10,7 +10,7 @@
defaultsort cpo
-subsection {* Type @{typ unit} is a pcpo *}
+subsection {* Continuous case function for unit type *}
definition
unit_when :: "'a \<rightarrow> unit \<rightarrow> 'a" where
@@ -22,150 +22,22 @@
lemma unit_when [simp]: "unit_when\<cdot>a\<cdot>u = a"
by (simp add: unit_when_def)
-subsection {* Continuous versions of constants *}
-
-definition
- cpair :: "'a \<rightarrow> 'b \<rightarrow> ('a * 'b)" -- {* continuous pairing *} where
- "cpair = (\<Lambda> x y. (x, y))"
-
-definition
- cfst :: "('a * 'b) \<rightarrow> 'a" where
- "cfst = (\<Lambda> p. fst p)"
-
-definition
- csnd :: "('a * 'b) \<rightarrow> 'b" where
- "csnd = (\<Lambda> p. snd p)"
+subsection {* Continuous version of split function *}
definition
csplit :: "('a \<rightarrow> 'b \<rightarrow> 'c) \<rightarrow> ('a * 'b) \<rightarrow> 'c" where
- "csplit = (\<Lambda> f p. f\<cdot>(cfst\<cdot>p)\<cdot>(csnd\<cdot>p))"
-
-syntax
- "_ctuple" :: "['a, args] \<Rightarrow> 'a * 'b" ("(1<_,/ _>)")
-
-syntax (xsymbols)
- "_ctuple" :: "['a, args] \<Rightarrow> 'a * 'b" ("(1\<langle>_,/ _\<rangle>)")
+ "csplit = (\<Lambda> f p. f\<cdot>(fst p)\<cdot>(snd p))"
translations
- "\<langle>x, y, z\<rangle>" == "\<langle>x, \<langle>y, z\<rangle>\<rangle>"
- "\<langle>x, y\<rangle>" == "CONST cpair\<cdot>x\<cdot>y"
-
-translations
- "\<Lambda>(CONST cpair\<cdot>x\<cdot>y). t" == "CONST csplit\<cdot>(\<Lambda> x y. t)"
- "\<Lambda>(CONST Pair x y). t" => "CONST csplit\<cdot>(\<Lambda> x y. t)"
+ "\<Lambda>(CONST Pair x y). t" == "CONST csplit\<cdot>(\<Lambda> x y. t)"
subsection {* Convert all lemmas to the continuous versions *}
-lemma cpair_eq_pair: "<x, y> = (x, y)"
-by (simp add: cpair_def cont_pair1 cont_pair2)
-
-lemma pair_eq_cpair: "(x, y) = <x, y>"
-by (simp add: cpair_def cont_pair1 cont_pair2)
-
-lemma inject_cpair: "<a,b> = <aa,ba> \<Longrightarrow> a = aa \<and> b = ba"
-by (simp add: cpair_eq_pair)
-
-lemma cpair_eq [iff]: "(<a, b> = <a', b'>) = (a = a' \<and> b = b')"
-by (simp add: cpair_eq_pair)
-
-lemma cpair_below [iff]: "(<a, b> \<sqsubseteq> <a', b'>) = (a \<sqsubseteq> a' \<and> b \<sqsubseteq> b')"
-by (simp add: cpair_eq_pair)
-
-lemma cpair_defined_iff [iff]: "(<x, y> = \<bottom>) = (x = \<bottom> \<and> y = \<bottom>)"
-by (simp add: cpair_eq_pair)
-
-lemma cpair_strict [simp]: "\<langle>\<bottom>, \<bottom>\<rangle> = \<bottom>"
-by simp
-
-lemma inst_cprod_pcpo2: "\<bottom> = <\<bottom>, \<bottom>>"
-by (rule cpair_strict [symmetric])
-
-lemma defined_cpair_rev:
- "<a,b> = \<bottom> \<Longrightarrow> a = \<bottom> \<and> b = \<bottom>"
-by simp
-
-lemma Exh_Cprod2: "\<exists>a b. z = <a, b>"
-by (simp add: cpair_eq_pair)
-
-lemma cprodE: "\<lbrakk>\<And>x y. p = <x, y> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
-by (cut_tac Exh_Cprod2, auto)
-
-lemma cfst_cpair [simp]: "cfst\<cdot><x, y> = x"
-by (simp add: cpair_eq_pair cfst_def)
-
-lemma csnd_cpair [simp]: "csnd\<cdot><x, y> = y"
-by (simp add: cpair_eq_pair csnd_def)
-
-lemma cfst_strict [simp]: "cfst\<cdot>\<bottom> = \<bottom>"
-by (simp add: cfst_def)
-
-lemma csnd_strict [simp]: "csnd\<cdot>\<bottom> = \<bottom>"
-by (simp add: csnd_def)
-
-lemma cpair_cfst_csnd: "\<langle>cfst\<cdot>p, csnd\<cdot>p\<rangle> = p"
-by (cases p rule: cprodE, simp)
-
-lemmas surjective_pairing_Cprod2 = cpair_cfst_csnd
-
-lemma below_cprod: "x \<sqsubseteq> y = (cfst\<cdot>x \<sqsubseteq> cfst\<cdot>y \<and> csnd\<cdot>x \<sqsubseteq> csnd\<cdot>y)"
-by (simp add: below_prod_def cfst_def csnd_def)
-
-lemma eq_cprod: "(x = y) = (cfst\<cdot>x = cfst\<cdot>y \<and> csnd\<cdot>x = csnd\<cdot>y)"
-by (auto simp add: po_eq_conv below_cprod)
-
-lemma cfst_below_iff: "cfst\<cdot>x \<sqsubseteq> y = x \<sqsubseteq> <y, csnd\<cdot>x>"
-by (simp add: below_cprod)
-
-lemma csnd_below_iff: "csnd\<cdot>x \<sqsubseteq> y = x \<sqsubseteq> <cfst\<cdot>x, y>"
-by (simp add: below_cprod)
-
-lemma compact_cfst: "compact x \<Longrightarrow> compact (cfst\<cdot>x)"
-by (rule compactI, simp add: cfst_below_iff)
-
-lemma compact_csnd: "compact x \<Longrightarrow> compact (csnd\<cdot>x)"
-by (rule compactI, simp add: csnd_below_iff)
-
-lemma compact_cpair: "\<lbrakk>compact x; compact y\<rbrakk> \<Longrightarrow> compact <x, y>"
-by (simp add: cpair_eq_pair)
-
-lemma compact_cpair_iff [simp]: "compact <x, y> = (compact x \<and> compact y)"
-by (simp add: cpair_eq_pair)
-
-lemma lub_cprod2:
- "chain S \<Longrightarrow> range S <<| <\<Squnion>i. cfst\<cdot>(S i), \<Squnion>i. csnd\<cdot>(S i)>"
-apply (simp add: cpair_eq_pair cfst_def csnd_def)
-apply (erule lub_cprod)
-done
-
-lemma thelub_cprod2:
- "chain S \<Longrightarrow> (\<Squnion>i. S i) = <\<Squnion>i. cfst\<cdot>(S i), \<Squnion>i. csnd\<cdot>(S i)>"
-by (rule lub_cprod2 [THEN thelubI])
-
lemma csplit1 [simp]: "csplit\<cdot>f\<cdot>\<bottom> = f\<cdot>\<bottom>\<cdot>\<bottom>"
by (simp add: csplit_def)
-lemma csplit2 [simp]: "csplit\<cdot>f\<cdot><x,y> = f\<cdot>x\<cdot>y"
+lemma csplit_Pair [simp]: "csplit\<cdot>f\<cdot>(x, y) = f\<cdot>x\<cdot>y"
by (simp add: csplit_def)
-lemma csplit3 [simp]: "csplit\<cdot>cpair\<cdot>z = z"
-by (simp add: csplit_def cpair_cfst_csnd)
-
-lemma csplit_Pair [simp]: "csplit\<cdot>f\<cdot>(x, y) = f\<cdot>x\<cdot>y"
-by (simp add: csplit_def cfst_def csnd_def)
-
-lemmas Cprod_rews = cfst_cpair csnd_cpair csplit2
-
-subsection {* Product type is a bifinite domain *}
-
-lemma approx_cpair [simp]:
- "approx i\<cdot>\<langle>x, y\<rangle> = \<langle>approx i\<cdot>x, approx i\<cdot>y\<rangle>"
-by (simp add: cpair_eq_pair)
-
-lemma cfst_approx: "cfst\<cdot>(approx i\<cdot>p) = approx i\<cdot>(cfst\<cdot>p)"
-by (cases p rule: cprodE, simp)
-
-lemma csnd_approx: "csnd\<cdot>(approx i\<cdot>p) = approx i\<cdot>(csnd\<cdot>p)"
-by (cases p rule: cprodE, simp)
-
end
--- a/src/HOLCF/Deflation.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Deflation.thy Thu Mar 25 17:56:31 2010 +0100
@@ -2,7 +2,7 @@
Author: Brian Huffman
*)
-header {* Continuous deflations and embedding-projection pairs *}
+header {* Continuous deflations and ep-pairs *}
theory Deflation
imports Cfun
--- a/src/HOLCF/Discrete.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Discrete.thy Thu Mar 25 17:56:31 2010 +0100
@@ -10,7 +10,7 @@
datatype 'a discr = Discr "'a :: type"
-subsection {* Type @{typ "'a discr"} is a discrete cpo *}
+subsection {* Discrete ordering *}
instantiation discr :: (type) below
begin
@@ -22,14 +22,14 @@
instance ..
end
+subsection {* Discrete cpo class instance *}
+
instance discr :: (type) discrete_cpo
by intro_classes (simp add: below_discr_def)
lemma discr_below_eq [iff]: "((x::('a::type)discr) << y) = (x = y)"
by simp (* FIXME: same discrete_cpo - remove? is [iff] important? *)
-subsection {* Type @{typ "'a discr"} is a cpo *}
-
lemma discr_chain0:
"!!S::nat=>('a::type)discr. chain S ==> S i = S 0"
apply (unfold chain_def)
@@ -60,7 +60,7 @@
apply (clarify, erule discr_chain0 [symmetric])
done
-subsection {* @{term undiscr} *}
+subsection {* \emph{undiscr} *}
definition
undiscr :: "('a::type)discr => 'a" where
--- a/src/HOLCF/FOCUS/Fstreams.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/FOCUS/Fstreams.thy Thu Mar 25 17:56:31 2010 +0100
@@ -324,15 +324,6 @@
lemma cpo_cont_lemma:
"[| monofun (f::'a::cpo => 'b::cpo); (!Y. chain Y --> f (lub(range Y)) << (LUB i. f (Y i))) |] ==> cont f"
-apply (rule monocontlub2cont, auto)
-apply (simp add: contlub_def, auto)
-apply (erule_tac x="Y" in allE, auto)
-apply (simp add: po_eq_conv)
-apply (frule cpo,auto)
-apply (frule is_lubD1)
-apply (frule ub2ub_monofun, auto)
-apply (drule thelubI, auto)
-apply (rule is_lub_thelub, auto)
-by (erule ch2ch_monofun, simp)
+by (erule contI2, simp)
end
--- a/src/HOLCF/Ffun.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Ffun.thy Thu Mar 25 17:56:31 2010 +0100
@@ -196,22 +196,13 @@
text {* the lub of a chain of continuous functions is continuous *}
-lemma contlub_lub_fun:
- "\<lbrakk>chain F; \<forall>i. cont (F i)\<rbrakk> \<Longrightarrow> contlub (\<Squnion>i. F i)"
-apply (rule contlubI)
-apply (simp add: thelub_fun)
-apply (simp add: cont2contlubE)
-apply (rule ex_lub)
-apply (erule ch2ch_fun)
-apply (simp add: ch2ch_cont)
-done
-
lemma cont_lub_fun:
"\<lbrakk>chain F; \<forall>i. cont (F i)\<rbrakk> \<Longrightarrow> cont (\<Squnion>i. F i)"
-apply (rule monocontlub2cont)
+apply (rule contI2)
apply (erule monofun_lub_fun)
apply (simp add: cont2mono)
-apply (erule (1) contlub_lub_fun)
+apply (simp add: thelub_fun cont2contlubE)
+apply (simp add: diag_lub ch2ch_fun ch2ch_cont)
done
lemma cont2cont_lub:
@@ -224,9 +215,8 @@
done
lemma cont2cont_fun: "cont f \<Longrightarrow> cont (\<lambda>x. f x y)"
-apply (rule monocontlub2cont)
+apply (rule contI2)
apply (erule cont2mono [THEN mono2mono_fun])
-apply (rule contlubI)
apply (simp add: cont2contlubE)
apply (simp add: thelub_fun ch2ch_cont)
done
@@ -242,14 +232,10 @@
lemma cont2cont_lambda [simp]:
assumes f: "\<And>y. cont (\<lambda>x. f x y)" shows "cont f"
-apply (subgoal_tac "monofun f")
-apply (rule monocontlub2cont)
-apply assumption
-apply (rule contlubI)
-apply (rule ext)
-apply (simp add: thelub_fun ch2ch_monofun)
-apply (erule cont2contlubE [OF f])
+apply (rule contI2)
apply (simp add: mono2mono_lambda cont2mono f)
+apply (rule below_fun_ext)
+apply (simp add: thelub_fun cont2contlubE [OF f])
done
text {* What D.A.Schmidt calls continuity of abstraction; never used here *}
@@ -272,23 +258,12 @@
apply (simp add: monofun_fun monofunE)
done
-lemma cont2contlub_app:
- "\<lbrakk>cont f; \<forall>x. cont (f x); cont t\<rbrakk> \<Longrightarrow> contlub (\<lambda>x. (f x) (t x))"
-apply (rule contlubI)
-apply (subgoal_tac "chain (\<lambda>i. f (Y i))")
-apply (subgoal_tac "chain (\<lambda>i. t (Y i))")
-apply (simp add: cont2contlubE thelub_fun)
-apply (rule diag_lub)
-apply (erule ch2ch_fun)
-apply (drule spec)
-apply (erule (1) ch2ch_cont)
-apply (erule (1) ch2ch_cont)
-apply (erule (1) ch2ch_cont)
-done
-
lemma cont2cont_app:
"\<lbrakk>cont f; \<forall>x. cont (f x); cont t\<rbrakk> \<Longrightarrow> cont (\<lambda>x. (f x) (t x))"
-by (blast intro: monocontlub2cont mono2mono_app cont2mono cont2contlub_app)
+apply (erule cont_apply [where t=t])
+apply (erule spec)
+apply (erule cont2cont_fun)
+done
lemmas cont2cont_app2 = cont2cont_app [rule_format]
--- a/src/HOLCF/Fix.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Fix.thy Thu Mar 25 17:56:31 2010 +0100
@@ -6,7 +6,7 @@
header {* Fixed point operator and admissibility *}
theory Fix
-imports Cfun Cprod
+imports Cfun
begin
defaultsort pcpo
@@ -203,31 +203,7 @@
by (simp add: fix_def2)
qed
-subsection {* Recursive let bindings *}
-
-definition
- CLetrec :: "('a \<rightarrow> 'a \<times> 'b) \<rightarrow> 'b" where
- "CLetrec = (\<Lambda> F. csnd\<cdot>(F\<cdot>(\<mu> x. cfst\<cdot>(F\<cdot>x))))"
-
-nonterminals
- recbinds recbindt recbind
-
-syntax
- "_recbind" :: "['a, 'a] \<Rightarrow> recbind" ("(2_ =/ _)" 10)
- "" :: "recbind \<Rightarrow> recbindt" ("_")
- "_recbindt" :: "[recbind, recbindt] \<Rightarrow> recbindt" ("_,/ _")
- "" :: "recbindt \<Rightarrow> recbinds" ("_")
- "_recbinds" :: "[recbindt, recbinds] \<Rightarrow> recbinds" ("_;/ _")
- "_Letrec" :: "[recbinds, 'a] \<Rightarrow> 'a" ("(Letrec (_)/ in (_))" 10)
-
-translations
- (recbindt) "x = a, \<langle>y,ys\<rangle> = \<langle>b,bs\<rangle>" == (recbindt) "\<langle>x,y,ys\<rangle> = \<langle>a,b,bs\<rangle>"
- (recbindt) "x = a, y = b" == (recbindt) "\<langle>x,y\<rangle> = \<langle>a,b\<rangle>"
-
-translations
- "_Letrec (_recbinds b bs) e" == "_Letrec b (_Letrec bs e)"
- "Letrec xs = a in \<langle>e,es\<rangle>" == "CONST CLetrec\<cdot>(\<Lambda> xs. \<langle>a,e,es\<rangle>)"
- "Letrec xs = a in e" == "CONST CLetrec\<cdot>(\<Lambda> xs. \<langle>a,e\<rangle>)"
+subsection {* Fixed-points on product types *}
text {*
Bekic's Theorem: Simultaneous fixed points over pairs
@@ -236,54 +212,31 @@
lemma fix_cprod:
"fix\<cdot>(F::'a \<times> 'b \<rightarrow> 'a \<times> 'b) =
- \<langle>\<mu> x. cfst\<cdot>(F\<cdot>\<langle>x, \<mu> y. csnd\<cdot>(F\<cdot>\<langle>x, y\<rangle>)\<rangle>),
- \<mu> y. csnd\<cdot>(F\<cdot>\<langle>\<mu> x. cfst\<cdot>(F\<cdot>\<langle>x, \<mu> y. csnd\<cdot>(F\<cdot>\<langle>x, y\<rangle>)\<rangle>), y\<rangle>)\<rangle>"
- (is "fix\<cdot>F = \<langle>?x, ?y\<rangle>")
+ (\<mu> x. fst (F\<cdot>(x, \<mu> y. snd (F\<cdot>(x, y)))),
+ \<mu> y. snd (F\<cdot>(\<mu> x. fst (F\<cdot>(x, \<mu> y. snd (F\<cdot>(x, y)))), y)))"
+ (is "fix\<cdot>F = (?x, ?y)")
proof (rule fix_eqI)
- have 1: "cfst\<cdot>(F\<cdot>\<langle>?x, ?y\<rangle>) = ?x"
+ have 1: "fst (F\<cdot>(?x, ?y)) = ?x"
by (rule trans [symmetric, OF fix_eq], simp)
- have 2: "csnd\<cdot>(F\<cdot>\<langle>?x, ?y\<rangle>) = ?y"
+ have 2: "snd (F\<cdot>(?x, ?y)) = ?y"
by (rule trans [symmetric, OF fix_eq], simp)
- from 1 2 show "F\<cdot>\<langle>?x, ?y\<rangle> = \<langle>?x, ?y\<rangle>" by (simp add: eq_cprod)
+ from 1 2 show "F\<cdot>(?x, ?y) = (?x, ?y)" by (simp add: Pair_fst_snd_eq)
next
fix z assume F_z: "F\<cdot>z = z"
- then obtain x y where z: "z = \<langle>x,y\<rangle>" by (rule_tac p=z in cprodE)
- from F_z z have F_x: "cfst\<cdot>(F\<cdot>\<langle>x, y\<rangle>) = x" by simp
- from F_z z have F_y: "csnd\<cdot>(F\<cdot>\<langle>x, y\<rangle>) = y" by simp
- let ?y1 = "\<mu> y. csnd\<cdot>(F\<cdot>\<langle>x, y\<rangle>)"
+ obtain x y where z: "z = (x,y)" by (rule prod.exhaust)
+ from F_z z have F_x: "fst (F\<cdot>(x, y)) = x" by simp
+ from F_z z have F_y: "snd (F\<cdot>(x, y)) = y" by simp
+ let ?y1 = "\<mu> y. snd (F\<cdot>(x, y))"
have "?y1 \<sqsubseteq> y" by (rule fix_least, simp add: F_y)
- hence "cfst\<cdot>(F\<cdot>\<langle>x, ?y1\<rangle>) \<sqsubseteq> cfst\<cdot>(F\<cdot>\<langle>x, y\<rangle>)" by (simp add: monofun_cfun)
- hence "cfst\<cdot>(F\<cdot>\<langle>x, ?y1\<rangle>) \<sqsubseteq> x" using F_x by simp
+ hence "fst (F\<cdot>(x, ?y1)) \<sqsubseteq> fst (F\<cdot>(x, y))"
+ by (simp add: fst_monofun monofun_cfun)
+ hence "fst (F\<cdot>(x, ?y1)) \<sqsubseteq> x" using F_x by simp
hence 1: "?x \<sqsubseteq> x" by (simp add: fix_least_below)
- hence "csnd\<cdot>(F\<cdot>\<langle>?x, y\<rangle>) \<sqsubseteq> csnd\<cdot>(F\<cdot>\<langle>x, y\<rangle>)" by (simp add: monofun_cfun)
- hence "csnd\<cdot>(F\<cdot>\<langle>?x, y\<rangle>) \<sqsubseteq> y" using F_y by simp
+ hence "snd (F\<cdot>(?x, y)) \<sqsubseteq> snd (F\<cdot>(x, y))"
+ by (simp add: snd_monofun monofun_cfun)
+ hence "snd (F\<cdot>(?x, y)) \<sqsubseteq> y" using F_y by simp
hence 2: "?y \<sqsubseteq> y" by (simp add: fix_least_below)
- show "\<langle>?x, ?y\<rangle> \<sqsubseteq> z" using z 1 2 by simp
+ show "(?x, ?y) \<sqsubseteq> z" using z 1 2 by simp
qed
-subsection {* Weak admissibility *}
-
-definition
- admw :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where
- "admw P = (\<forall>F. (\<forall>n. P (iterate n\<cdot>F\<cdot>\<bottom>)) \<longrightarrow> P (\<Squnion>i. iterate i\<cdot>F\<cdot>\<bottom>))"
-
-text {* an admissible formula is also weak admissible *}
-
-lemma adm_impl_admw: "adm P \<Longrightarrow> admw P"
-apply (unfold admw_def)
-apply (intro strip)
-apply (erule admD)
-apply (rule chain_iterate)
-apply (erule spec)
-done
-
-text {* computational induction for weak admissible formulae *}
-
-lemma wfix_ind: "\<lbrakk>admw P; \<forall>n. P (iterate n\<cdot>F\<cdot>\<bottom>)\<rbrakk> \<Longrightarrow> P (fix\<cdot>F)"
-by (simp add: fix_def2 admw_def)
-
-lemma def_wfix_ind:
- "\<lbrakk>f \<equiv> fix\<cdot>F; admw P; \<forall>n. P (iterate n\<cdot>F\<cdot>\<bottom>)\<rbrakk> \<Longrightarrow> P f"
-by (simp, rule wfix_ind)
-
end
--- a/src/HOLCF/Fixrec.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Fixrec.thy Thu Mar 25 17:56:31 2010 +0100
@@ -5,7 +5,7 @@
header "Package for defining recursive functions in HOLCF"
theory Fixrec
-imports Sprod Ssum Up One Tr Fix
+imports Cprod Sprod Ssum Up One Tr Fix
uses
("Tools/holcf_library.ML")
("Tools/fixrec.ML")
@@ -276,13 +276,13 @@
definition
cpair_pat :: "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a \<times> 'b, 'c \<times> 'd) pat" where
- "cpair_pat p1 p2 = (\<Lambda>\<langle>x, y\<rangle>.
- bind\<cdot>(p1\<cdot>x)\<cdot>(\<Lambda> a. bind\<cdot>(p2\<cdot>y)\<cdot>(\<Lambda> b. return\<cdot>\<langle>a, b\<rangle>)))"
+ "cpair_pat p1 p2 = (\<Lambda>(x, y).
+ bind\<cdot>(p1\<cdot>x)\<cdot>(\<Lambda> a. bind\<cdot>(p2\<cdot>y)\<cdot>(\<Lambda> b. return\<cdot>(a, b))))"
definition
spair_pat ::
"('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a::pcpo \<otimes> 'b::pcpo, 'c \<times> 'd) pat" where
- "spair_pat p1 p2 = (\<Lambda>(:x, y:). cpair_pat p1 p2\<cdot>\<langle>x, y\<rangle>)"
+ "spair_pat p1 p2 = (\<Lambda>(:x, y:). cpair_pat p1 p2\<cdot>(x, y))"
definition
sinl_pat :: "('a, 'c) pat \<Rightarrow> ('a::pcpo \<oplus> 'b::pcpo, 'c) pat" where
@@ -310,7 +310,7 @@
text {* Parse translations (patterns) *}
translations
- "_pat (XCONST cpair\<cdot>x\<cdot>y)" => "CONST cpair_pat (_pat x) (_pat y)"
+ "_pat (XCONST Pair x y)" => "CONST cpair_pat (_pat x) (_pat y)"
"_pat (XCONST spair\<cdot>x\<cdot>y)" => "CONST spair_pat (_pat x) (_pat y)"
"_pat (XCONST sinl\<cdot>x)" => "CONST sinl_pat (_pat x)"
"_pat (XCONST sinr\<cdot>x)" => "CONST sinr_pat (_pat x)"
@@ -321,12 +321,12 @@
text {* CONST version is also needed for constructors with special syntax *}
translations
- "_pat (CONST cpair\<cdot>x\<cdot>y)" => "CONST cpair_pat (_pat x) (_pat y)"
+ "_pat (CONST Pair x y)" => "CONST cpair_pat (_pat x) (_pat y)"
"_pat (CONST spair\<cdot>x\<cdot>y)" => "CONST spair_pat (_pat x) (_pat y)"
text {* Parse translations (variables) *}
translations
- "_variable (XCONST cpair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
+ "_variable (XCONST Pair x y) r" => "_variable (_args x y) r"
"_variable (XCONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
"_variable (XCONST sinl\<cdot>x) r" => "_variable x r"
"_variable (XCONST sinr\<cdot>x) r" => "_variable x r"
@@ -336,12 +336,12 @@
"_variable (XCONST ONE) r" => "_variable _noargs r"
translations
- "_variable (CONST cpair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
+ "_variable (CONST Pair x y) r" => "_variable (_args x y) r"
"_variable (CONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
text {* Print translations *}
translations
- "CONST cpair\<cdot>(_match p1 v1)\<cdot>(_match p2 v2)"
+ "CONST Pair (_match p1 v1) (_match p2 v2)"
<= "_match (CONST cpair_pat p1 p2) (_args v1 v2)"
"CONST spair\<cdot>(_match p1 v1)\<cdot>(_match p2 v2)"
<= "_match (CONST spair_pat p1 p2) (_args v1 v2)"
@@ -353,20 +353,20 @@
"CONST ONE" <= "_match (CONST ONE_pat) _noargs"
lemma cpair_pat1:
- "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>\<langle>x, y\<rangle> = \<bottom>"
+ "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = \<bottom>"
apply (simp add: branch_def cpair_pat_def)
apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
done
lemma cpair_pat2:
- "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>\<langle>x, y\<rangle> = fail"
+ "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = fail"
apply (simp add: branch_def cpair_pat_def)
apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
done
lemma cpair_pat3:
"branch p\<cdot>r\<cdot>x = return\<cdot>s \<Longrightarrow>
- branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>\<langle>x, y\<rangle> = branch q\<cdot>s\<cdot>y"
+ branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = branch q\<cdot>s\<cdot>y"
apply (simp add: branch_def cpair_pat_def)
apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
apply (rule_tac p="q\<cdot>y" in maybeE, simp_all)
@@ -379,7 +379,7 @@
"branch (spair_pat p1 p2)\<cdot>r\<cdot>\<bottom> = \<bottom>"
"\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk>
\<Longrightarrow> branch (spair_pat p1 p2)\<cdot>r\<cdot>(:x, y:) =
- branch (cpair_pat p1 p2)\<cdot>r\<cdot>\<langle>x, y\<rangle>"
+ branch (cpair_pat p1 p2)\<cdot>r\<cdot>(x, y)"
by (simp_all add: branch_def spair_pat_def)
lemma sinl_pat [simp]:
@@ -419,17 +419,13 @@
subsection {* Wildcards, as-patterns, and lazy patterns *}
-syntax
- "_as_pat" :: "[idt, 'a] \<Rightarrow> 'a" (infixr "\<as>" 10)
- "_lazy_pat" :: "'a \<Rightarrow> 'a" ("\<lazy> _" [1000] 1000)
-
definition
wild_pat :: "'a \<rightarrow> unit maybe" where
"wild_pat = (\<Lambda> x. return\<cdot>())"
definition
as_pat :: "('a \<rightarrow> 'b maybe) \<Rightarrow> 'a \<rightarrow> ('a \<times> 'b) maybe" where
- "as_pat p = (\<Lambda> x. bind\<cdot>(p\<cdot>x)\<cdot>(\<Lambda> a. return\<cdot>\<langle>x, a\<rangle>))"
+ "as_pat p = (\<Lambda> x. bind\<cdot>(p\<cdot>x)\<cdot>(\<Lambda> a. return\<cdot>(x, a)))"
definition
lazy_pat :: "('a \<rightarrow> 'b::pcpo maybe) \<Rightarrow> ('a \<rightarrow> 'b maybe)" where
@@ -438,24 +434,14 @@
text {* Parse translations (patterns) *}
translations
"_pat _" => "CONST wild_pat"
- "_pat (_as_pat x y)" => "CONST as_pat (_pat y)"
- "_pat (_lazy_pat x)" => "CONST lazy_pat (_pat x)"
text {* Parse translations (variables) *}
translations
"_variable _ r" => "_variable _noargs r"
- "_variable (_as_pat x y) r" => "_variable (_args x y) r"
- "_variable (_lazy_pat x) r" => "_variable x r"
text {* Print translations *}
translations
"_" <= "_match (CONST wild_pat) _noargs"
- "_as_pat x (_match p v)" <= "_match (CONST as_pat p) (_args (_variable x) v)"
- "_lazy_pat (_match p v)" <= "_match (CONST lazy_pat p) v"
-
-text {* Lazy patterns in lambda abstractions *}
-translations
- "_cabs (_lazy_pat p) r" == "CONST Fixrec.cases oo (_Case1 (_lazy_pat p) r)"
lemma wild_pat [simp]: "branch wild_pat\<cdot>(unit_when\<cdot>r)\<cdot>x = return\<cdot>r"
by (simp add: branch_def wild_pat_def)
@@ -530,7 +516,6 @@
by (simp_all add: match_UU_def)
lemma match_cpair_simps [simp]:
- "match_cpair\<cdot>\<langle>x, y\<rangle>\<cdot>k = k\<cdot>x\<cdot>y"
"match_cpair\<cdot>(x, y)\<cdot>k = k\<cdot>x\<cdot>y"
by (simp_all add: match_cpair_def)
@@ -616,7 +601,6 @@
(@{const_name sinl}, @{const_name match_sinl}),
(@{const_name sinr}, @{const_name match_sinr}),
(@{const_name spair}, @{const_name match_spair}),
- (@{const_name cpair}, @{const_name match_cpair}),
(@{const_name Pair}, @{const_name match_cpair}),
(@{const_name ONE}, @{const_name match_ONE}),
(@{const_name TT}, @{const_name match_TT}),
--- a/src/HOLCF/HOLCF.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/HOLCF.thy Thu Mar 25 17:56:31 2010 +0100
@@ -10,20 +10,10 @@
Domain
Powerdomains
Sum_Cpo
-uses
- "holcf_logic.ML"
- "Tools/adm_tac.ML"
begin
defaultsort pcpo
-declaration {* fn _ =>
- Simplifier.map_ss (fn simpset => simpset addSolver
- (mk_solver' "adm_tac" (fn ss =>
- Adm.adm_tac (Simplifier.the_context ss)
- (cut_facts_tac (Simplifier.prems_of_ss ss) THEN' cont_tacRs ss))));
-*}
-
text {* Legacy theorem names *}
lemmas sq_ord_less_eq_trans = below_eq_trans
@@ -60,10 +50,6 @@
lemmas not_Iup_less = not_Iup_below
lemmas Iup_less = Iup_below
lemmas up_less = up_below
-lemmas cpair_less = cpair_below
-lemmas less_cprod = below_cprod
-lemmas cfst_less_iff = cfst_below_iff
-lemmas csnd_less_iff = csnd_below_iff
lemmas Def_inject_less_eq = Def_below_Def
lemmas Def_less_is_eq = Def_below_iff
lemmas spair_less_iff = spair_below_iff
--- a/src/HOLCF/IOA/meta_theory/Seq.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/IOA/meta_theory/Seq.thy Thu Mar 25 17:56:31 2010 +0100
@@ -175,7 +175,7 @@
szip_nil: "szip$nil$y=nil"
| szip_cons_nil: "x~=UU ==> szip$(x##xs)$nil=UU"
| szip_cons:
- "[| x~=UU; y~=UU|] ==> szip$(x##xs)$(y##ys) = <x,y>##szip$xs$ys"
+ "[| x~=UU; y~=UU|] ==> szip$(x##xs)$(y##ys) = (x,y)##szip$xs$ys"
lemma szip_UU1 [simp]: "szip$UU$y=UU"
by fixrec_simp
--- a/src/HOLCF/IOA/meta_theory/Sequence.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/IOA/meta_theory/Sequence.thy Thu Mar 25 17:56:31 2010 +0100
@@ -814,9 +814,8 @@
lemma nForall_HDFilter:
"~Forall P y --> (? x. HD$(Filter (%a. ~P a)$y) = Def x)"
-(* Pay attention: is only admissible with chain-finite package to be added to
- adm test *)
-apply (rule_tac x="y" in Seq_induct)
+unfolding not_Undef_is_Def [symmetric]
+apply (induct y rule: Seq_induct)
apply (simp add: Forall_def sforall_def)
apply simp_all
done
--- a/src/HOLCF/IsaMakefile Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/IsaMakefile Thu Mar 25 17:56:31 2010 +0100
@@ -61,7 +61,6 @@
Universal.thy \
UpperPD.thy \
Up.thy \
- Tools/adm_tac.ML \
Tools/cont_consts.ML \
Tools/cont_proc.ML \
Tools/holcf_library.ML \
@@ -75,7 +74,6 @@
Tools/fixrec.ML \
Tools/pcpodef.ML \
Tools/repdef.ML \
- holcf_logic.ML \
document/root.tex
@$(ISABELLE_TOOL) usedir -b -g true -r $(OUT)/HOL HOLCF
@@ -103,6 +101,7 @@
ex/Fixrec_ex.thy \
ex/Focus_ex.thy \
ex/Hoare.thy \
+ ex/Letrec.thy \
ex/Loop.thy \
ex/New_Domain.thy \
ex/Powerdomain_ex.thy \
--- a/src/HOLCF/Lift.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Lift.thy Thu Mar 25 17:56:31 2010 +0100
@@ -170,28 +170,6 @@
lemma flift2_defined_iff [simp]: "(flift2 f\<cdot>x = \<bottom>) = (x = \<bottom>)"
by (cases x, simp_all)
-text {*
- \medskip Extension of @{text cont_tac} and installation of simplifier.
-*}
-
-lemmas cont_lemmas_ext =
- cont2cont_flift1 cont2cont_lift_case cont2cont_lambda
- cont_Rep_CFun_app cont_Rep_CFun_app_app cont_if
-
-ML {*
-local
- val cont_lemmas2 = thms "cont_lemmas1" @ thms "cont_lemmas_ext";
- val flift1_def = thm "flift1_def";
-in
-
-fun cont_tac i = resolve_tac cont_lemmas2 i;
-fun cont_tacR i = REPEAT (cont_tac i);
-
-fun cont_tacRs ss i =
- simp_tac ss i THEN
- REPEAT (cont_tac i)
-end;
-*}
subsection {* Lifted countable types are bifinite *}
--- a/src/HOLCF/LowerPD.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/LowerPD.thy Thu Mar 25 17:56:31 2010 +0100
@@ -498,13 +498,15 @@
lemma ep_pair_lower_map: "ep_pair e p \<Longrightarrow> ep_pair (lower_map\<cdot>e) (lower_map\<cdot>p)"
apply default
apply (induct_tac x rule: lower_pd_induct, simp_all add: ep_pair.e_inverse)
-apply (induct_tac y rule: lower_pd_induct, simp_all add: ep_pair.e_p_below monofun_cfun)
+apply (induct_tac y rule: lower_pd_induct)
+apply (simp_all add: ep_pair.e_p_below monofun_cfun)
done
lemma deflation_lower_map: "deflation d \<Longrightarrow> deflation (lower_map\<cdot>d)"
apply default
apply (induct_tac x rule: lower_pd_induct, simp_all add: deflation.idem)
-apply (induct_tac x rule: lower_pd_induct, simp_all add: deflation.below monofun_cfun)
+apply (induct_tac x rule: lower_pd_induct)
+apply (simp_all add: deflation.below monofun_cfun)
done
end
--- a/src/HOLCF/Pcpodef.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Pcpodef.thy Thu Mar 25 17:56:31 2010 +0100
@@ -129,7 +129,7 @@
thus "\<exists>x. range S <<| x" ..
qed
-subsubsection {* Continuity of @{term Rep} and @{term Abs} *}
+subsubsection {* Continuity of \emph{Rep} and \emph{Abs} *}
text {* For any sub-cpo, the @{term Rep} function is continuous. *}
@@ -229,7 +229,7 @@
shows "OFCLASS('b, pcpo_class)"
by (rule typedef_pcpo_generic [OF type below UU_in_A], rule minimal)
-subsubsection {* Strictness of @{term Rep} and @{term Abs} *}
+subsubsection {* Strictness of \emph{Rep} and \emph{Abs} *}
text {*
For a sub-pcpo where @{term \<bottom>} is a member of the defining
--- a/src/HOLCF/Product_Cpo.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Product_Cpo.thy Thu Mar 25 17:56:31 2010 +0100
@@ -10,7 +10,7 @@
defaultsort cpo
-subsection {* Type @{typ unit} is a pcpo *}
+subsection {* Unit type is a pcpo *}
instantiation unit :: below
begin
@@ -58,7 +58,7 @@
by (fast intro: below_trans)
qed
-subsection {* Monotonicity of @{text "(_,_)"}, @{term fst}, @{term snd} *}
+subsection {* Monotonicity of \emph{Pair}, \emph{fst}, \emph{snd} *}
lemma prod_belowI: "\<lbrakk>fst p \<sqsubseteq> fst q; snd p \<sqsubseteq> snd q\<rbrakk> \<Longrightarrow> p \<sqsubseteq> q"
unfolding below_prod_def by simp
@@ -84,6 +84,12 @@
text {* @{term fst} and @{term snd} are monotone *}
+lemma fst_monofun: "x \<sqsubseteq> y \<Longrightarrow> fst x \<sqsubseteq> fst y"
+unfolding below_prod_def by simp
+
+lemma snd_monofun: "x \<sqsubseteq> y \<Longrightarrow> snd x \<sqsubseteq> snd y"
+unfolding below_prod_def by simp
+
lemma monofun_fst: "monofun fst"
by (simp add: monofun_def below_prod_def)
@@ -187,7 +193,7 @@
lemma split_strict [simp]: "split f \<bottom> = f \<bottom> \<bottom>"
unfolding split_def by simp
-subsection {* Continuity of @{text "(_,_)"}, @{term fst}, @{term snd} *}
+subsection {* Continuity of \emph{Pair}, \emph{fst}, \emph{snd} *}
lemma cont_pair1: "cont (\<lambda>x. (x, y))"
apply (rule contI)
@@ -203,26 +209,16 @@
apply (erule cpo_lubI)
done
-lemma contlub_fst: "contlub fst"
-apply (rule contlubI)
+lemma cont_fst: "cont fst"
+apply (rule contI)
apply (simp add: thelub_cprod)
-done
-
-lemma contlub_snd: "contlub snd"
-apply (rule contlubI)
-apply (simp add: thelub_cprod)
-done
-
-lemma cont_fst: "cont fst"
-apply (rule monocontlub2cont)
-apply (rule monofun_fst)
-apply (rule contlub_fst)
+apply (erule cpo_lubI [OF ch2ch_fst])
done
lemma cont_snd: "cont snd"
-apply (rule monocontlub2cont)
-apply (rule monofun_snd)
-apply (rule contlub_snd)
+apply (rule contI)
+apply (simp add: thelub_cprod)
+apply (erule cpo_lubI [OF ch2ch_snd])
done
lemma cont2cont_Pair [cont2cont]:
--- a/src/HOLCF/Representable.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Representable.thy Thu Mar 25 17:56:31 2010 +0100
@@ -35,7 +35,7 @@
lemmas prj_strict = rep.p_strict
-subsection {* Making @{term rep} the default class *}
+subsection {* Making \emph{rep} the default class *}
text {*
From now on, free type variables are assumed to be in class
@@ -342,7 +342,7 @@
use "Tools/repdef.ML"
-subsection {* Instances of class @{text rep} *}
+subsection {* Instances of class \emph{rep} *}
subsubsection {* Universal Domain *}
@@ -517,7 +517,8 @@
lemma cast_TypeRep_fun2:
assumes f: "\<And>a b. finite_deflation a \<Longrightarrow> finite_deflation b \<Longrightarrow>
finite_deflation (f\<cdot>a\<cdot>b)"
- shows "cast\<cdot>(TypeRep_fun2 f\<cdot>A\<cdot>B) = udom_emb oo f\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj"
+ shows "cast\<cdot>(TypeRep_fun2 f\<cdot>A\<cdot>B) =
+ udom_emb oo f\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj"
proof -
have 1: "\<And>a b. finite_deflation
(udom_emb oo f\<cdot>(Rep_fin_defl a)\<cdot>(Rep_fin_defl b) oo udom_prj)"
--- a/src/HOLCF/Sprod.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Sprod.thy Thu Mar 25 17:56:31 2010 +0100
@@ -88,7 +88,7 @@
"\<lbrakk>P \<bottom>; \<And>x y. \<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> P (:x, y:)\<rbrakk> \<Longrightarrow> P x"
by (cases x, simp_all)
-subsection {* Properties of @{term spair} *}
+subsection {* Properties of \emph{spair} *}
lemma spair_strict1 [simp]: "(:\<bottom>, y:) = \<bottom>"
by (simp add: Rep_Sprod_simps strictify_conv_if)
@@ -134,7 +134,7 @@
lemma sprodE2: "(\<And>x y. p = (:x, y:) \<Longrightarrow> Q) \<Longrightarrow> Q"
by (cases p, simp only: inst_sprod_pcpo2, simp)
-subsection {* Properties of @{term sfst} and @{term ssnd} *}
+subsection {* Properties of \emph{sfst} and \emph{ssnd} *}
lemma sfst_strict [simp]: "sfst\<cdot>\<bottom> = \<bottom>"
by (simp add: sfst_def cont_Rep_Sprod Rep_Sprod_strict)
@@ -208,7 +208,7 @@
apply simp
done
-subsection {* Properties of @{term ssplit} *}
+subsection {* Properties of \emph{ssplit} *}
lemma ssplit1 [simp]: "ssplit\<cdot>f\<cdot>\<bottom> = \<bottom>"
by (simp add: ssplit_def)
--- a/src/HOLCF/Ssum.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Ssum.thy Thu Mar 25 17:56:31 2010 +0100
@@ -58,7 +58,7 @@
lemma Rep_Ssum_sinr: "Rep_Ssum (sinr\<cdot>b) = (strictify\<cdot>(\<Lambda> _. FF)\<cdot>b, \<bottom>, b)"
by (simp add: sinr_Abs_Ssum Abs_Ssum_inverse sinr_Ssum)
-subsection {* Properties of @{term sinl} and @{term sinr} *}
+subsection {* Properties of \emph{sinl} and \emph{sinr} *}
text {* Ordering *}
--- a/src/HOLCF/Sum_Cpo.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Sum_Cpo.thy Thu Mar 25 17:56:31 2010 +0100
@@ -8,7 +8,7 @@
imports Bifinite
begin
-subsection {* Ordering on type @{typ "'a + 'b"} *}
+subsection {* Ordering on sum type *}
instantiation "+" :: (below, below) below
begin
@@ -128,7 +128,7 @@
apply (erule cpo_lubI)
done
-subsection {* Continuity of @{term Inl}, @{term Inr}, @{term sum_case} *}
+subsection {* Continuity of \emph{Inl}, \emph{Inr}, and case function *}
lemma cont_Inl: "cont Inl"
by (intro contI is_lub_Inl cpo_lubI)
--- a/src/HOLCF/Tools/Domain/domain_axioms.ML Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Tools/Domain/domain_axioms.ML Thu Mar 25 17:56:31 2010 +0100
@@ -54,8 +54,8 @@
val abs_iso_bind = Binding.qualified true "abs_iso" dbind;
val rep_iso_bind = Binding.qualified true "rep_iso" dbind;
- val (abs_iso_thm, thy) = Drule.add_axiom (abs_iso_bind, abs_iso_eqn) thy;
- val (rep_iso_thm, thy) = Drule.add_axiom (rep_iso_bind, rep_iso_eqn) thy;
+ val (abs_iso_thm, thy) = Specification.axiom ((abs_iso_bind, []), abs_iso_eqn) thy;
+ val (rep_iso_thm, thy) = Specification.axiom ((rep_iso_bind, []), rep_iso_eqn) thy;
val result =
{
@@ -63,8 +63,8 @@
repT = rhsT,
abs_const = abs_const,
rep_const = rep_const,
- abs_inverse = abs_iso_thm,
- rep_inverse = rep_iso_thm
+ abs_inverse = Drule.export_without_context abs_iso_thm,
+ rep_inverse = Drule.export_without_context rep_iso_thm
};
in
(result, thy)
@@ -83,9 +83,9 @@
val lub_take_bind = Binding.qualified true "lub_take" dbind;
- val (lub_take_thm, thy) = Drule.add_axiom (lub_take_bind, lub_take_eqn) thy;
+ val (lub_take_thm, thy) = Specification.axiom ((lub_take_bind, []), lub_take_eqn) thy;
in
- (lub_take_thm, thy)
+ (Drule.export_without_context lub_take_thm, thy)
end;
fun add_axioms
--- a/src/HOLCF/Tools/Domain/domain_library.ML Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Tools/Domain/domain_library.ML Thu Mar 25 17:56:31 2010 +0100
@@ -212,6 +212,7 @@
fun con_app2 con f args = list_ccomb(%%:con,map f args);
fun con_app con = con_app2 con %#;
fun prj _ _ x ( _::[]) _ = x
+ | prj _ _ _ [] _ = raise Fail "Domain_Library.prj: empty list"
| prj f1 _ x (_::y::ys) 0 = f1 x y
| prj f1 f2 x (y:: ys) j = prj f1 f2 (f2 x y) ys (j-1);
fun proj x = prj (fn S => K(%%:"fst" $S)) (fn S => K(%%:"snd" $S)) x;
--- a/src/HOLCF/Tools/adm_tac.ML Thu Mar 25 17:55:55 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,160 +0,0 @@
-(* Title: HOLCF/Tools/adm_tac.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Admissibility tactic.
-
-Checks whether adm_subst theorem is applicable to the current proof
-state:
-
- cont t ==> adm P ==> adm (%x. P (t x))
-
-"t" is instantiated with a term of chain-finite type, so that
-adm_chfin can be applied:
-
- adm (P::'a::{chfin,pcpo} => bool)
-*)
-
-signature ADM =
-sig
- val adm_tac: Proof.context -> (int -> tactic) -> int -> tactic
-end;
-
-structure Adm : ADM =
-struct
-
-
-(*** find_subterms t 0 []
- returns lists of terms with the following properties:
- 1. all terms in the list are disjoint subterms of t
- 2. all terms contain the variable which is bound at level 0
- 3. all occurences of the variable which is bound at level 0
- are "covered" by a term in the list
- a list of integers is associated with every term which describes
- the "path" leading to the subterm (required for instantiation of
- the adm_subst theorem (see functions mk_term, inst_adm_subst_thm))
-***)
-
-fun find_subterms (Bound i) lev path =
- if i = lev then [[(Bound 0, path)]]
- else []
- | find_subterms (t as (Abs (_, _, t2))) lev path =
- if filter (fn x => x <= lev) (add_loose_bnos (t, 0, [])) = [lev]
- then
- [(incr_bv (~lev, 0, t), path)] ::
- (find_subterms t2 (lev+1) (0::path))
- else find_subterms t2 (lev+1) (0::path)
- | find_subterms (t as (t1 $ t2)) lev path =
- let val ts1 = find_subterms t1 lev (0::path);
- val ts2 = find_subterms t2 lev (1::path);
- fun combine [] y = []
- | combine (x::xs) ys = map (fn z => x @ z) ys @ combine xs ys
- in
- (if filter (fn x => x <= lev) (add_loose_bnos (t, 0, [])) = [lev]
- then [[(incr_bv (~lev, 0, t), path)]]
- else []) @
- (if ts1 = [] then ts2
- else if ts2 = [] then ts1
- else combine ts1 ts2)
- end
- | find_subterms _ _ _ = [];
-
-
-(*** make term for instantiation of predicate "P" in adm_subst theorem ***)
-
-fun make_term t path paths lev =
- if member (op =) paths path then Bound lev
- else case t of
- (Abs (s, T, t1)) => Abs (s, T, make_term t1 (0::path) paths (lev+1))
- | (t1 $ t2) => (make_term t1 (0::path) paths lev) $
- (make_term t2 (1::path) paths lev)
- | t1 => t1;
-
-
-(*** check whether all terms in list are equal ***)
-
-fun eq_terms [] = true
- | eq_terms (ts as (t, _) :: _) = forall (fn (t2, _) => t2 aconv t) ts;
-
-
-(*** check whether type of terms in list is chain finite ***)
-
-fun is_chfin thy T params ((t, _)::_) =
- let val parTs = map snd (rev params)
- in Sign.of_sort thy (fastype_of1 (T::parTs, t), @{sort "{chfin,pcpo}"}) end;
-
-
-(*** try to prove that terms in list are continuous
- if successful, add continuity theorem to list l ***)
-
-fun prove_cont ctxt tac s T prems params (ts as ((t, _)::_)) l = (* FIXME proper context *)
- let val parTs = map snd (rev params);
- val contT = (T --> (fastype_of1 (T::parTs, t))) --> HOLogic.boolT;
- fun mk_all [] t = t
- | mk_all ((a,T)::Ts) t = Term.all T $ (Abs (a, T, mk_all Ts t));
- val t = HOLogic.mk_Trueprop (Const (@{const_name cont}, contT) $ Abs (s, T, t));
- val t' = mk_all params (Logic.list_implies (prems, t));
- val thm = Goal.prove ctxt [] [] t' (K (tac 1));
- in (ts, thm)::l end
- handle ERROR _ => l;
-
-
-(*** instantiation of adm_subst theorem (a bit tricky) ***)
-
-fun inst_adm_subst_thm state i params s T subt t paths =
- let
- val thy = Thm.theory_of_thm state;
- val j = Thm.maxidx_of state + 1;
- val parTs = map snd (rev params);
- val rule = Thm.lift_rule (Thm.cprem_of state i) @{thm adm_subst};
- val types = the o fst (Drule.types_sorts rule);
- val tT = types ("t", j);
- val PT = types ("P", j);
- fun mk_abs [] t = t
- | mk_abs ((a,T)::Ts) t = Abs (a, T, mk_abs Ts t);
- val tt = cterm_of thy (mk_abs (params @ [(s, T)]) subt);
- val Pt = cterm_of thy (mk_abs (params @ [(s, fastype_of1 (T::parTs, subt))])
- (make_term t [] paths 0));
- val tye = Sign.typ_match thy (tT, #T (rep_cterm tt)) Vartab.empty;
- val tye' = Sign.typ_match thy (PT, #T (rep_cterm Pt)) tye;
- val ctye = map (fn (ixn, (S, T)) =>
- (ctyp_of thy (TVar (ixn, S)), ctyp_of thy T)) (Vartab.dest tye');
- val tv = cterm_of thy (Var (("t", j), Envir.subst_type tye' tT));
- val Pv = cterm_of thy (Var (("P", j), Envir.subst_type tye' PT));
- val rule' = instantiate (ctye, [(tv, tt), (Pv, Pt)]) rule
- in rule' end;
-
-
-(*** the admissibility tactic ***)
-
-fun try_dest_adm (Const _ $ (Const (@{const_name adm}, _) $ Abs abs)) = SOME abs
- | try_dest_adm _ = NONE;
-
-fun adm_tac ctxt tac i state = (i, state) |-> SUBGOAL (fn (goali, _) =>
- (case try_dest_adm (Logic.strip_assums_concl goali) of
- NONE => no_tac
- | SOME (s, T, t) =>
- let
- val thy = ProofContext.theory_of ctxt;
- val prems = Logic.strip_assums_hyp goali;
- val params = Logic.strip_params goali;
- val ts = find_subterms t 0 [];
- val ts' = filter eq_terms ts;
- val ts'' = filter (is_chfin thy T params) ts';
- val thms = fold (prove_cont ctxt tac s T prems params) ts'' [];
- in
- (case thms of
- ((ts as ((t', _)::_), cont_thm) :: _) =>
- let
- val paths = map snd ts;
- val rule = inst_adm_subst_thm state i params s T t' t paths;
- in
- compose_tac (false, rule, 2) i THEN
- resolve_tac [cont_thm] i THEN
- REPEAT (assume_tac i) THEN
- resolve_tac [@{thm adm_chfin}] i
- end
- | [] => no_tac)
- end));
-
-end;
-
--- a/src/HOLCF/Tools/fixrec.ML Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Tools/fixrec.ML Thu Mar 25 17:56:31 2010 +0100
@@ -275,12 +275,16 @@
(* this is the pattern-matching compiler function *)
fun compile_pats match_name eqs =
let
- val (((n::names),(a::arities)),mats) =
+ val ((names, arities), mats) =
apfst ListPair.unzip (ListPair.unzip (map (match_eq match_name) eqs));
- val cname = if forall (fn x => n=x) names then n
- else fixrec_err "all equations in block must define the same function";
- val arity = if forall (fn x => a=x) arities then a
- else fixrec_err "all equations in block must have the same arity";
+ val cname =
+ case distinct (op =) names of
+ [n] => n
+ | _ => fixrec_err "all equations in block must define the same function";
+ val arity =
+ case distinct (op =) arities of
+ [a] => a
+ | _ => fixrec_err "all equations in block must have the same arity";
val rhs = fatbar arity mats;
in
mk_trp (cname === rhs)
@@ -311,7 +315,8 @@
else HOLogic.dest_Trueprop (Logic.strip_imp_concl t);
fun tac (t, i) =
let
- val Const (c, T) = head_of (chead_of (fst (HOLogic.dest_eq (concl t))));
+ val (c, T) =
+ (dest_Const o head_of o chead_of o fst o HOLogic.dest_eq o concl) t;
val unfold_thm = the (Symtab.lookup tab c);
val rule = unfold_thm RS @{thm ssubst_lhs};
in
--- a/src/HOLCF/Tools/holcf_library.ML Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Tools/holcf_library.ML Thu Mar 25 17:56:31 2010 +0100
@@ -62,7 +62,6 @@
(*** Continuous function space ***)
-(* ->> is taken from holcf_logic.ML *)
fun mk_cfunT (T, U) = Type(@{type_name cfun}, [T, U]);
val (op ->>) = mk_cfunT;
@@ -228,7 +227,7 @@
in
(v::vs, mk_ssumT (T, U))
end
- fun inj [] = error "mk_sinjects: empty list"
+ fun inj [] = raise Fail "mk_sinjects: empty list"
| inj ((t, T)::[]) = ([t], T)
| inj ((t, T)::ts) = combine (t, T) (inj ts);
in
--- a/src/HOLCF/Tools/pcpodef.ML Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Tools/pcpodef.ML Thu Mar 25 17:56:31 2010 +0100
@@ -87,29 +87,32 @@
val cpo_thms = map (Thm.transfer thy) [type_definition, below_def, admissible'];
val (full_tname, Ts) = dest_Type newT;
val lhs_sorts = map (snd o dest_TFree) Ts;
- val thy2 =
- thy
- |> AxClass.prove_arity (full_tname, lhs_sorts, @{sort cpo})
- (Tactic.rtac (@{thm typedef_cpo} OF cpo_thms) 1);
+ val tac = Tactic.rtac (@{thm typedef_cpo} OF cpo_thms) 1;
+ val thy = AxClass.prove_arity (full_tname, lhs_sorts, @{sort cpo}) tac thy;
(* transfer thms so that they will know about the new cpo instance *)
- val cpo_thms' = map (Thm.transfer thy2) cpo_thms;
+ val cpo_thms' = map (Thm.transfer thy) cpo_thms;
fun make thm = Drule.export_without_context (thm OF cpo_thms');
- val ([adm, cont_Rep, cont_Abs, lub, thelub, compact], thy3) =
- thy2
+ val cont_Rep = make @{thm typedef_cont_Rep};
+ val cont_Abs = make @{thm typedef_cont_Abs};
+ val lub = make @{thm typedef_lub};
+ val thelub = make @{thm typedef_thelub};
+ val compact = make @{thm typedef_compact};
+ val (_, thy) =
+ thy
|> Sign.add_path (Binding.name_of name)
|> PureThy.add_thms
- ([((Binding.prefix_name "adm_" name, admissible'), []),
- ((Binding.prefix_name "cont_" Rep_name, make @{thm typedef_cont_Rep}), []),
- ((Binding.prefix_name "cont_" Abs_name, make @{thm typedef_cont_Abs}), []),
- ((Binding.prefix_name "lub_" name, make @{thm typedef_lub}), []),
- ((Binding.prefix_name "thelub_" name, make @{thm typedef_thelub}), []),
- ((Binding.prefix_name "compact_" name, make @{thm typedef_compact}), [])])
- ||> Sign.restore_naming thy2;
+ ([((Binding.prefix_name "adm_" name, admissible'), []),
+ ((Binding.prefix_name "cont_" Rep_name, cont_Rep ), []),
+ ((Binding.prefix_name "cont_" Abs_name, cont_Abs ), []),
+ ((Binding.prefix_name "lub_" name, lub ), []),
+ ((Binding.prefix_name "thelub_" name, thelub ), []),
+ ((Binding.prefix_name "compact_" name, compact ), [])])
+ ||> Sign.parent_path;
val cpo_info : cpo_info =
{ below_def = below_def, adm = admissible', cont_Rep = cont_Rep,
cont_Abs = cont_Abs, lub = lub, thelub = thelub, compact = compact };
in
- (cpo_info, thy3)
+ (cpo_info, thy)
end;
fun prove_pcpo
@@ -127,29 +130,33 @@
val pcpo_thms = map (Thm.transfer thy) [type_definition, below_def, UU_mem'];
val (full_tname, Ts) = dest_Type newT;
val lhs_sorts = map (snd o dest_TFree) Ts;
- val thy2 = thy
- |> AxClass.prove_arity (full_tname, lhs_sorts, @{sort pcpo})
- (Tactic.rtac (@{thm typedef_pcpo} OF pcpo_thms) 1);
- val pcpo_thms' = map (Thm.transfer thy2) pcpo_thms;
+ val tac = Tactic.rtac (@{thm typedef_pcpo} OF pcpo_thms) 1;
+ val thy = AxClass.prove_arity (full_tname, lhs_sorts, @{sort pcpo}) tac thy;
+ val pcpo_thms' = map (Thm.transfer thy) pcpo_thms;
fun make thm = Drule.export_without_context (thm OF pcpo_thms');
- val ([Rep_strict, Abs_strict, Rep_strict_iff, Abs_strict_iff,
- Rep_defined, Abs_defined], thy3) =
- thy2
+ val Rep_strict = make @{thm typedef_Rep_strict};
+ val Abs_strict = make @{thm typedef_Abs_strict};
+ val Rep_strict_iff = make @{thm typedef_Rep_strict_iff};
+ val Abs_strict_iff = make @{thm typedef_Abs_strict_iff};
+ val Rep_defined = make @{thm typedef_Rep_defined};
+ val Abs_defined = make @{thm typedef_Abs_defined};
+ val (_, thy) =
+ thy
|> Sign.add_path (Binding.name_of name)
|> PureThy.add_thms
- ([((Binding.suffix_name "_strict" Rep_name, make @{thm typedef_Rep_strict}), []),
- ((Binding.suffix_name "_strict" Abs_name, make @{thm typedef_Abs_strict}), []),
- ((Binding.suffix_name "_strict_iff" Rep_name, make @{thm typedef_Rep_strict_iff}), []),
- ((Binding.suffix_name "_strict_iff" Abs_name, make @{thm typedef_Abs_strict_iff}), []),
- ((Binding.suffix_name "_defined" Rep_name, make @{thm typedef_Rep_defined}), []),
- ((Binding.suffix_name "_defined" Abs_name, make @{thm typedef_Abs_defined}), [])])
- ||> Sign.restore_naming thy2;
+ ([((Binding.suffix_name "_strict" Rep_name, Rep_strict), []),
+ ((Binding.suffix_name "_strict" Abs_name, Abs_strict), []),
+ ((Binding.suffix_name "_strict_iff" Rep_name, Rep_strict_iff), []),
+ ((Binding.suffix_name "_strict_iff" Abs_name, Abs_strict_iff), []),
+ ((Binding.suffix_name "_defined" Rep_name, Rep_defined), []),
+ ((Binding.suffix_name "_defined" Abs_name, Abs_defined), [])])
+ ||> Sign.parent_path;
val pcpo_info =
{ Rep_strict = Rep_strict, Abs_strict = Abs_strict,
Rep_strict_iff = Rep_strict_iff, Abs_strict_iff = Abs_strict_iff,
Rep_defined = Rep_defined, Abs_defined = Abs_defined };
in
- (pcpo_info, thy3)
+ (pcpo_info, thy)
end;
(* prepare_cpodef *)
@@ -319,7 +326,8 @@
val args = map (apsnd (prep_constraint ctxt)) raw_args;
val (goal1, goal2, make_result) =
prepare_cpodef prep_term def name (b, args, mx) set opt_morphs thy;
- fun after_qed [[th1, th2]] = ProofContext.theory (snd o make_result th1 th2);
+ fun after_qed [[th1, th2]] = ProofContext.theory (snd o make_result th1 th2)
+ | after_qed _ = raise Fail "cpodef_proof";
in Proof.theorem_i NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end;
fun gen_pcpodef_proof prep_term prep_constraint
@@ -329,7 +337,8 @@
val args = map (apsnd (prep_constraint ctxt)) raw_args;
val (goal1, goal2, make_result) =
prepare_pcpodef prep_term def name (b, args, mx) set opt_morphs thy;
- fun after_qed [[th1, th2]] = ProofContext.theory (snd o make_result th1 th2);
+ fun after_qed [[th1, th2]] = ProofContext.theory (snd o make_result th1 th2)
+ | after_qed _ = raise Fail "pcpodef_proof";
in Proof.theorem_i NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end;
in
--- a/src/HOLCF/Tools/repdef.ML Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Tools/repdef.ML Thu Mar 25 17:56:31 2010 +0100
@@ -68,14 +68,14 @@
thy |> Theory.copy |> Theory_Target.init NONE
|> Typedecl.predeclare_constraints (tname, raw_args, mx);
val defl = prep_term tmp_lthy raw_defl;
- val tmp_lthy' = tmp_lthy |> Variable.declare_constraints defl;
+ val tmp_lthy = tmp_lthy |> Variable.declare_constraints defl;
val deflT = Term.fastype_of defl;
val _ = if deflT = @{typ "udom alg_defl"} then ()
else error ("Not type udom alg_defl: " ^ quote (Syntax.string_of_typ tmp_lthy deflT));
(*lhs*)
- val lhs_tfrees = map (fn (a, _) => (a, ProofContext.default_sort tmp_lthy' (a, ~1))) raw_args;
+ val lhs_tfrees = map (fn (a, _) => (a, ProofContext.default_sort tmp_lthy (a, ~1))) raw_args;
val lhs_sorts = map snd lhs_tfrees;
val full_tname = Sign.full_name thy tname;
val newT = Type (full_tname, map TFree lhs_tfrees);
@@ -91,7 +91,7 @@
(*pcpodef*)
val tac1 = rtac @{thm CollectI} 1 THEN rtac @{thm bottom_in_deflation} 1;
val tac2 = rtac @{thm adm_mem_Collect_in_deflation} 1;
- val ((info, cpo_info, pcpo_info), thy2) = thy
+ val ((info, cpo_info, pcpo_info), thy) = thy
|> Pcpodef.add_pcpodef def (SOME name) typ set (SOME morphs) (tac1, tac2);
(*definitions*)
@@ -105,44 +105,49 @@
--> alg_deflT udomT --> natT --> (newT ->> newT));
val approx_eqn = Logic.mk_equals (approx_const newT,
repdef_approx_const $ Rep_const $ Abs_const $ defl);
+ val name_def = Binding.suffix_name "_def" name;
+ val emb_bind = (Binding.prefix_name "emb_" name_def, []);
+ val prj_bind = (Binding.prefix_name "prj_" name_def, []);
+ val approx_bind = (Binding.prefix_name "approx_" name_def, []);
(*instantiate class rep*)
- val name_def = Binding.suffix_name "_def" name;
- val ([emb_ldef, prj_ldef, approx_ldef], lthy3) = thy2
- |> Theory_Target.instantiation ([full_tname], lhs_tfrees, @{sort rep})
- |> fold_map Specification.definition
- [ (NONE, ((Binding.prefix_name "emb_" name_def, []), emb_eqn))
- , (NONE, ((Binding.prefix_name "prj_" name_def, []), prj_eqn))
- , (NONE, ((Binding.prefix_name "approx_" name_def, []), approx_eqn)) ]
- |>> map (snd o snd);
- val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy3);
- val [emb_def, prj_def, approx_def] =
- ProofContext.export lthy3 ctxt_thy [emb_ldef, prj_ldef, approx_ldef];
+ val lthy = thy
+ |> Theory_Target.instantiation ([full_tname], lhs_tfrees, @{sort rep});
+ val ((_, (_, emb_ldef)), lthy) =
+ Specification.definition (NONE, (emb_bind, emb_eqn)) lthy;
+ val ((_, (_, prj_ldef)), lthy) =
+ Specification.definition (NONE, (prj_bind, prj_eqn)) lthy;
+ val ((_, (_, approx_ldef)), lthy) =
+ Specification.definition (NONE, (approx_bind, approx_eqn)) lthy;
+ val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy);
+ val emb_def = singleton (ProofContext.export lthy ctxt_thy) emb_ldef;
+ val prj_def = singleton (ProofContext.export lthy ctxt_thy) prj_ldef;
+ val approx_def = singleton (ProofContext.export lthy ctxt_thy) approx_ldef;
val type_definition_thm =
MetaSimplifier.rewrite_rule
(the_list (#set_def info))
(#type_definition info);
val typedef_thms =
[type_definition_thm, #below_def cpo_info, emb_def, prj_def, approx_def];
- val thy4 = lthy3
+ val thy = lthy
|> Class.prove_instantiation_instance
(K (Tactic.rtac (@{thm typedef_rep_class} OF typedef_thms) 1))
|> Local_Theory.exit_global;
(*other theorems*)
- val typedef_thms' = map (Thm.transfer thy4)
+ val typedef_thms' = map (Thm.transfer thy)
[type_definition_thm, #below_def cpo_info, emb_def, prj_def];
- val ([REP_thm], thy5) = thy4
+ val (REP_thm, thy) = thy
|> Sign.add_path (Binding.name_of name)
- |> PureThy.add_thms
- [((Binding.prefix_name "REP_" name,
- Drule.export_without_context (@{thm typedef_REP} OF typedef_thms')), [])]
- ||> Sign.restore_naming thy4;
+ |> PureThy.add_thm
+ ((Binding.prefix_name "REP_" name,
+ Drule.export_without_context (@{thm typedef_REP} OF typedef_thms')), [])
+ ||> Sign.restore_naming thy;
val rep_info =
{ emb_def = emb_def, prj_def = prj_def, approx_def = approx_def, REP = REP_thm };
in
- ((info, cpo_info, pcpo_info, rep_info), thy5)
+ ((info, cpo_info, pcpo_info, rep_info), thy)
end
handle ERROR msg =>
cat_error msg ("The error(s) above occurred in repdef " ^ quote (Binding.str_of name));
--- a/src/HOLCF/Universal.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Universal.thy Thu Mar 25 17:56:31 2010 +0100
@@ -187,7 +187,7 @@
apply simp
done
-subsubsection {* Take function for @{typ ubasis} *}
+subsubsection {* Take function for \emph{ubasis} *}
definition
ubasis_take :: "nat \<Rightarrow> ubasis \<Rightarrow> ubasis"
@@ -338,7 +338,7 @@
by (rule udom.completion_approx_eq_principal)
-subsection {* Universality of @{typ udom} *}
+subsection {* Universality of \emph{udom} *}
defaultsort bifinite
@@ -816,7 +816,7 @@
place
sub
-subsubsection {* EP-pair from any bifinite domain into @{typ udom} *}
+subsubsection {* EP-pair from any bifinite domain into \emph{udom} *}
definition
udom_emb :: "'a::bifinite \<rightarrow> udom"
--- a/src/HOLCF/Up.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/Up.thy Thu Mar 25 17:56:31 2010 +0100
@@ -169,7 +169,7 @@
lemma inst_up_pcpo: "\<bottom> = Ibottom"
by (rule minimal_up [THEN UU_I, symmetric])
-subsection {* Continuity of @{term Iup} and @{term Ifup} *}
+subsection {* Continuity of \emph{Iup} and \emph{Ifup} *}
text {* continuity for @{term Iup} *}
--- a/src/HOLCF/UpperPD.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/UpperPD.thy Thu Mar 25 17:56:31 2010 +0100
@@ -493,13 +493,15 @@
lemma ep_pair_upper_map: "ep_pair e p \<Longrightarrow> ep_pair (upper_map\<cdot>e) (upper_map\<cdot>p)"
apply default
apply (induct_tac x rule: upper_pd_induct, simp_all add: ep_pair.e_inverse)
-apply (induct_tac y rule: upper_pd_induct, simp_all add: ep_pair.e_p_below monofun_cfun)
+apply (induct_tac y rule: upper_pd_induct)
+apply (simp_all add: ep_pair.e_p_below monofun_cfun)
done
lemma deflation_upper_map: "deflation d \<Longrightarrow> deflation (upper_map\<cdot>d)"
apply default
apply (induct_tac x rule: upper_pd_induct, simp_all add: deflation.idem)
-apply (induct_tac x rule: upper_pd_induct, simp_all add: deflation.below monofun_cfun)
+apply (induct_tac x rule: upper_pd_induct)
+apply (simp_all add: deflation.below monofun_cfun)
done
end
--- a/src/HOLCF/document/root.tex Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/document/root.tex Thu Mar 25 17:56:31 2010 +0100
@@ -21,7 +21,7 @@
\tableofcontents
\begin{center}
- \includegraphics[scale=0.5]{session_graph}
+ \includegraphics[scale=0.45]{session_graph}
\end{center}
\newpage
--- a/src/HOLCF/ex/Dagstuhl.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/ex/Dagstuhl.thy Thu Mar 25 17:56:31 2010 +0100
@@ -72,7 +72,7 @@
apply (unfold YYS_def)
apply (rule fix_least)
apply (subst beta_cfun)
- apply (tactic "cont_tacR 1")
+ apply simp
apply (simp add: YS_def2 [symmetric])
done
--- a/src/HOLCF/ex/Fixrec_ex.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/ex/Fixrec_ex.thy Thu Mar 25 17:56:31 2010 +0100
@@ -13,7 +13,7 @@
text {*
Fixrec patterns can mention any constructor defined by the domain
package, as well as any of the following built-in constructors:
- cpair, spair, sinl, sinr, up, ONE, TT, FF.
+ Pair, spair, sinl, sinr, up, ONE, TT, FF.
*}
text {* Typical usage is with lazy constructors. *}
@@ -50,7 +50,7 @@
fixrec
lzip :: "'a llist \<rightarrow> 'b llist \<rightarrow> ('a \<times> 'b) llist"
where
- "lzip\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>(lCons\<cdot>y\<cdot>ys) = lCons\<cdot><x,y>\<cdot>(lzip\<cdot>xs\<cdot>ys)"
+ "lzip\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>(lCons\<cdot>y\<cdot>ys) = lCons\<cdot>(x, y)\<cdot>(lzip\<cdot>xs\<cdot>ys)"
| "lzip\<cdot>lNil\<cdot>lNil = lNil"
text {* @{text fixrec_simp} is useful for producing strictness theorems. *}
@@ -118,7 +118,7 @@
fixrec (permissive)
lzip2 :: "'a llist \<rightarrow> 'b llist \<rightarrow> ('a \<times> 'b) llist"
where
- "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>(lCons\<cdot>y\<cdot>ys) = lCons\<cdot><x,y>\<cdot>(lzip\<cdot>xs\<cdot>ys)"
+ "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>(lCons\<cdot>y\<cdot>ys) = lCons\<cdot>(x, y)\<cdot>(lzip\<cdot>xs\<cdot>ys)"
| "lzip2\<cdot>xs\<cdot>ys = lNil"
text {*
@@ -130,7 +130,7 @@
text {* Simp rules can be generated later using @{text fixrec_simp}. *}
lemma lzip2_simps [simp]:
- "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>(lCons\<cdot>y\<cdot>ys) = lCons\<cdot><x,y>\<cdot>(lzip\<cdot>xs\<cdot>ys)"
+ "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>(lCons\<cdot>y\<cdot>ys) = lCons\<cdot>(x, y)\<cdot>(lzip\<cdot>xs\<cdot>ys)"
"lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>lNil = lNil"
"lzip2\<cdot>lNil\<cdot>(lCons\<cdot>y\<cdot>ys) = lNil"
"lzip2\<cdot>lNil\<cdot>lNil = lNil"
--- a/src/HOLCF/ex/Focus_ex.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/ex/Focus_ex.thy Thu Mar 25 17:56:31 2010 +0100
@@ -30,7 +30,7 @@
input channel x:'b
output channel y:'c
is network
- <y,z> = f$<x,z>
+ (y,z) = f$(x,z)
end network
end g
@@ -47,7 +47,7 @@
'c stream * ('b,'c) tc stream) => bool
is_f f = !i1 i2 o1 o2.
- f$<i1,i2> = <o1,o2> --> Rf(i1,i2,o1,o2)
+ f$(i1,i2) = (o1,o2) --> Rf(i1,i2,o1,o2)
Specification of agent g is translated to predicate is_g which uses
predicate is_net_g
@@ -56,8 +56,8 @@
'b stream => 'c stream => bool
is_net_g f x y =
- ? z. <y,z> = f$<x,z> &
- !oy hz. <oy,hz> = f$<x,hz> --> z << hz
+ ? z. (y,z) = f$(x,z) &
+ !oy hz. (oy,hz) = f$(x,hz) --> z << hz
is_g :: ('b stream -> 'c stream) => bool
@@ -84,7 +84,7 @@
def_g g =
(? f. is_f f &
- g = (LAM x. cfst$(f$<x,fix$(LAM k.csnd$(f$<x,k>))>)) )
+ g = (LAM x. fst (f$(x,fix$(LAM k. snd (f$(x,k)))))) )
Now we prove:
@@ -110,14 +110,14 @@
definition
is_f :: "('b stream * ('b,'c) tc stream -> 'c stream * ('b,'c) tc stream) => bool" where
- "is_f f = (!i1 i2 o1 o2. f$<i1,i2> = <o1,o2> --> Rf(i1,i2,o1,o2))"
+ "is_f f = (!i1 i2 o1 o2. f$(i1,i2) = (o1,o2) --> Rf(i1,i2,o1,o2))"
definition
is_net_g :: "('b stream *('b,'c) tc stream -> 'c stream * ('b,'c) tc stream) =>
'b stream => 'c stream => bool" where
"is_net_g f x y == (? z.
- <y,z> = f$<x,z> &
- (!oy hz. <oy,hz> = f$<x,hz> --> z << hz))"
+ (y,z) = f$(x,z) &
+ (!oy hz. (oy,hz) = f$(x,hz) --> z << hz))"
definition
is_g :: "('b stream -> 'c stream) => bool" where
@@ -125,27 +125,27 @@
definition
def_g :: "('b stream -> 'c stream) => bool" where
- "def_g g == (? f. is_f f & g = (LAM x. cfst$(f$<x,fix$(LAM k. csnd$(f$<x,k>))>)))"
+ "def_g g == (? f. is_f f & g = (LAM x. fst (f$(x,fix$(LAM k. snd (f$(x,k)))))))"
(* first some logical trading *)
lemma lemma1:
"is_g(g) =
- (? f. is_f(f) & (!x.(? z. <g$x,z> = f$<x,z> &
- (! w y. <y,w> = f$<x,w> --> z << w))))"
+ (? f. is_f(f) & (!x.(? z. (g$x,z) = f$(x,z) &
+ (! w y. (y,w) = f$(x,w) --> z << w))))"
apply (simp add: is_g_def is_net_g_def)
apply fast
done
lemma lemma2:
-"(? f. is_f(f) & (!x. (? z. <g$x,z> = f$<x,z> &
- (!w y. <y,w> = f$<x,w> --> z << w))))
+"(? f. is_f(f) & (!x. (? z. (g$x,z) = f$(x,z) &
+ (!w y. (y,w) = f$(x,w) --> z << w))))
=
(? f. is_f(f) & (!x. ? z.
- g$x = cfst$(f$<x,z>) &
- z = csnd$(f$<x,z>) &
- (! w y. <y,w> = f$<x,w> --> z << w)))"
+ g$x = fst (f$(x,z)) &
+ z = snd (f$(x,z)) &
+ (! w y. (y,w) = f$(x,w) --> z << w)))"
apply (rule iffI)
apply (erule exE)
apply (rule_tac x = "f" in exI)
@@ -174,11 +174,9 @@
apply (erule conjE)+
apply (rule conjI)
prefer 2 apply (assumption)
-apply (rule trans)
-apply (rule_tac [2] surjective_pairing_Cprod2)
-apply (erule subst)
-apply (erule subst)
-apply (rule refl)
+apply (rule prod_eqI)
+apply simp
+apply simp
done
lemma lemma3: "def_g(g) --> is_g(g)"
@@ -189,12 +187,10 @@
apply (erule conjE)+
apply (erule conjI)
apply (intro strip)
-apply (rule_tac x = "fix$ (LAM k. csnd$ (f$<x,k>))" in exI)
+apply (rule_tac x = "fix$ (LAM k. snd (f$(x,k)))" in exI)
apply (rule conjI)
apply (tactic "asm_simp_tac HOLCF_ss 1")
- apply (rule trans)
- apply (rule_tac [2] surjective_pairing_Cprod2)
- apply (rule cfun_arg_cong)
+ apply (rule prod_eqI, simp, simp)
apply (rule trans)
apply (rule fix_eq)
apply (simp (no_asm))
@@ -219,20 +215,17 @@
apply (erule_tac x = "x" in allE)
apply (erule exE)
apply (erule conjE)+
-apply (subgoal_tac "fix$ (LAM k. csnd$ (f$<x, k>)) = z")
+apply (subgoal_tac "fix$ (LAM k. snd (f$(x, k))) = z")
apply simp
-apply (subgoal_tac "! w y. f$<x, w> = <y, w> --> z << w")
+apply (subgoal_tac "! w y. f$(x, w) = (y, w) --> z << w")
apply (rule fix_eqI)
apply simp
(*apply (rule allI)*)
(*apply (tactic "simp_tac HOLCF_ss 1")*)
(*apply (intro strip)*)
-apply (subgoal_tac "f$<x, za> = <cfst$ (f$<x,za>) ,za>")
+apply (subgoal_tac "f$(x, za) = (fst (f$(x,za)) ,za)")
apply fast
-apply (rule trans)
-apply (rule surjective_pairing_Cprod2 [symmetric])
-apply (rule cfun_arg_cong)
-apply simp
+apply (rule prod_eqI, simp, simp)
apply (intro strip)
apply (erule allE)+
apply (erule mp)
--- a/src/HOLCF/ex/Hoare.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/ex/Hoare.thy Thu Mar 25 17:56:31 2010 +0100
@@ -232,12 +232,8 @@
lemma fernpass_lemma: "(ALL k. b1$(iterate k$g$x)=TT) ==> ALL k. p$(iterate k$g$x) = UU"
apply (rule p_def [THEN eq_reflection, THEN def_fix_ind])
-apply (rule adm_all)
-apply (rule adm_eq)
-apply (tactic "cont_tacR 1")
-apply (rule allI)
-apply (subst Rep_CFun_strict1)
-apply (rule refl)
+apply simp
+apply simp
apply (simp (no_asm))
apply (rule allI)
apply (rule_tac s = "TT" and t = "b1$ (iterate k$g$x) " in ssubst)
@@ -256,12 +252,8 @@
lemma hoare_lemma17: "(ALL k. b1$(iterate k$g$x)=TT) ==> ALL k. q$(iterate k$g$x) = UU"
apply (rule q_def [THEN eq_reflection, THEN def_fix_ind])
-apply (rule adm_all)
-apply (rule adm_eq)
-apply (tactic "cont_tacR 1")
-apply (rule allI)
-apply (subst Rep_CFun_strict1)
-apply (rule refl)
+apply simp
+apply simp
apply (rule allI)
apply (simp (no_asm))
apply (rule_tac s = "TT" and t = "b1$ (iterate k$g$x) " in ssubst)
@@ -285,12 +277,8 @@
lemma hoare_lemma20: "(ALL y. b1$(y::'a)=TT) ==> ALL k. q$(iterate k$g$(x::'a)) = UU"
apply (rule q_def [THEN eq_reflection, THEN def_fix_ind])
-apply (rule adm_all)
-apply (rule adm_eq)
-apply (tactic "cont_tacR 1")
-apply (rule allI)
-apply (subst Rep_CFun_strict1)
-apply (rule refl)
+apply simp
+apply simp
apply (rule allI)
apply (simp (no_asm))
apply (rule_tac s = "TT" and t = "b1$ (iterate k$g$ (x::'a))" in ssubst)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/ex/Letrec.thy Thu Mar 25 17:56:31 2010 +0100
@@ -0,0 +1,37 @@
+(* Title: HOLCF/ex/Letrec.thy
+ Author: Brian Huffman
+*)
+
+header {* Recursive let bindings *}
+
+theory Letrec
+imports HOLCF
+begin
+
+defaultsort pcpo
+
+definition
+ CLetrec :: "('a \<rightarrow> 'a \<times> 'b) \<rightarrow> 'b" where
+ "CLetrec = (\<Lambda> F. snd (F\<cdot>(\<mu> x. fst (F\<cdot>x))))"
+
+nonterminals
+ recbinds recbindt recbind
+
+syntax
+ "_recbind" :: "['a, 'a] \<Rightarrow> recbind" ("(2_ =/ _)" 10)
+ "" :: "recbind \<Rightarrow> recbindt" ("_")
+ "_recbindt" :: "[recbind, recbindt] \<Rightarrow> recbindt" ("_,/ _")
+ "" :: "recbindt \<Rightarrow> recbinds" ("_")
+ "_recbinds" :: "[recbindt, recbinds] \<Rightarrow> recbinds" ("_;/ _")
+ "_Letrec" :: "[recbinds, 'a] \<Rightarrow> 'a" ("(Letrec (_)/ in (_))" 10)
+
+translations
+ (recbindt) "x = a, (y,ys) = (b,bs)" == (recbindt) "(x,y,ys) = (a,b,bs)"
+ (recbindt) "x = a, y = b" == (recbindt) "(x,y) = (a,b)"
+
+translations
+ "_Letrec (_recbinds b bs) e" == "_Letrec b (_Letrec bs e)"
+ "Letrec xs = a in (e,es)" == "CONST CLetrec\<cdot>(\<Lambda> xs. (a,e,es))"
+ "Letrec xs = a in e" == "CONST CLetrec\<cdot>(\<Lambda> xs. (a,e))"
+
+end
--- a/src/HOLCF/ex/Loop.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/ex/Loop.thy Thu Mar 25 17:56:31 2010 +0100
@@ -137,10 +137,8 @@
ALL m. while$b$g$(iterate m$(step$b$g)$x)=UU"
apply (simplesubst while_def2)
apply (rule fix_ind)
-apply (rule adm_all)
-apply (rule adm_eq)
-apply (tactic "cont_tacR 1")
-apply (simp (no_asm))
+apply simp
+apply simp
apply (rule allI)
apply (simp (no_asm))
apply (rule_tac p = "b$ (iterate m$ (step$b$g) $x) " in trE)
--- a/src/HOLCF/ex/Powerdomain_ex.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/ex/Powerdomain_ex.thy Thu Mar 25 17:56:31 2010 +0100
@@ -28,16 +28,16 @@
definition
r1 :: "(int lift \<times> 'a) \<rightarrow> (int lift \<times> 'a) \<rightarrow> tr convex_pd" where
- "r1 = (\<Lambda> \<langle>x,_\<rangle> \<langle>y,_\<rangle>. case compare\<cdot>x\<cdot>y of
+ "r1 = (\<Lambda> (x,_) (y,_). case compare\<cdot>x\<cdot>y of
LT \<Rightarrow> {TT}\<natural> |
EQ \<Rightarrow> {TT, FF}\<natural> |
GT \<Rightarrow> {FF}\<natural>)"
definition
r2 :: "(int lift \<times> 'a) \<rightarrow> (int lift \<times> 'a) \<rightarrow> tr convex_pd" where
- "r2 = (\<Lambda> \<langle>x,_\<rangle> \<langle>y,_\<rangle>. {is_le\<cdot>x\<cdot>y, is_less\<cdot>x\<cdot>y}\<natural>)"
+ "r2 = (\<Lambda> (x,_) (y,_). {is_le\<cdot>x\<cdot>y, is_less\<cdot>x\<cdot>y}\<natural>)"
-lemma r1_r2: "r1\<cdot>\<langle>x,a\<rangle>\<cdot>\<langle>y,b\<rangle> = (r2\<cdot>\<langle>x,a\<rangle>\<cdot>\<langle>y,b\<rangle> :: tr convex_pd)"
+lemma r1_r2: "r1\<cdot>(x,a)\<cdot>(y,b) = (r2\<cdot>(x,a)\<cdot>(y,b) :: tr convex_pd)"
apply (simp add: r1_def r2_def)
apply (simp add: is_le_def is_less_def)
apply (cases "compare\<cdot>x\<cdot>y")
@@ -57,8 +57,8 @@
mirror_Leaf: "mirror\<cdot>(Leaf\<cdot>a) = Leaf\<cdot>a"
| mirror_Node: "mirror\<cdot>(Node\<cdot>l\<cdot>r) = Node\<cdot>(mirror\<cdot>r)\<cdot>(mirror\<cdot>l)"
-fixpat
- mirror_strict [simp]: "mirror\<cdot>\<bottom>"
+lemma mirror_strict [simp]: "mirror\<cdot>\<bottom> = \<bottom>"
+by fixrec_simp
fixrec
pick :: "'a tree \<rightarrow> 'a convex_pd"
@@ -66,8 +66,8 @@
pick_Leaf: "pick\<cdot>(Leaf\<cdot>a) = {a}\<natural>"
| pick_Node: "pick\<cdot>(Node\<cdot>l\<cdot>r) = pick\<cdot>l +\<natural> pick\<cdot>r"
-fixpat
- pick_strict [simp]: "pick\<cdot>\<bottom>"
+lemma pick_strict [simp]: "pick\<cdot>\<bottom> = \<bottom>"
+by fixrec_simp
lemma pick_mirror: "pick\<cdot>(mirror\<cdot>t) = pick\<cdot>t"
by (induct t) (simp_all add: convex_plus_ac)
--- a/src/HOLCF/ex/ROOT.ML Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/ex/ROOT.ML Thu Mar 25 17:56:31 2010 +0100
@@ -5,5 +5,6 @@
use_thys ["Dnat", "Stream", "Dagstuhl", "Focus_ex", "Fix2", "Hoare",
"Loop", "Fixrec_ex", "Powerdomain_ex", "Domain_ex", "Domain_Proofs",
+ "Letrec",
"Strict_Fun",
"New_Domain"];
--- a/src/HOLCF/ex/Stream.thy Thu Mar 25 17:55:55 2010 +0100
+++ b/src/HOLCF/ex/Stream.thy Thu Mar 25 17:56:31 2010 +0100
@@ -920,11 +920,8 @@
apply (simp add: chain_def,auto)
by (rule monofun_cfun_arg,simp)
-lemma contlub_scons: "contlub (%x. a && x)"
-by (simp add: contlub_Rep_CFun2)
-
lemma contlub_scons_lemma: "chain S ==> (LUB i. a && S i) = a && (LUB i. S i)"
-by (rule contlubE [OF contlub_Rep_CFun2, symmetric])
+by (rule cont2contlubE [OF cont_Rep_CFun2, symmetric])
lemma finite_lub_sconc: "chain Y ==> (stream_finite x) ==>
(LUB i. x ooo Y i) = (x ooo (LUB i. Y i))"
@@ -940,9 +937,6 @@
apply (drule finite_lub_sconc,auto simp add: slen_infinite)
done
-lemma contlub_sconc: "contlub (%y. x ooo y)"
-by (rule cont_sconc [THEN cont2contlub])
-
lemma monofun_sconc: "monofun (%y. x ooo y)"
by (simp add: monofun_def sconc_mono)
--- a/src/HOLCF/holcf_logic.ML Thu Mar 25 17:55:55 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,44 +0,0 @@
-(* Title: HOLCF/holcf_logic.ML
- Author: David von Oheimb
-
-Abstract syntax operations for HOLCF.
-*)
-
-infixr 6 ->>;
-
-signature HOLCF_LOGIC =
-sig
- val mk_btyp: string -> typ * typ -> typ
- val pcpoS: sort
- val mk_TFree: string -> typ
- val cfun_arrow: string
- val ->> : typ * typ -> typ
- val mk_ssumT: typ * typ -> typ
- val mk_sprodT: typ * typ -> typ
- val mk_uT: typ -> typ
- val trT: typ
- val oneT: typ
-end;
-
-structure HOLCFLogic: HOLCF_LOGIC =
-struct
-
-(* sort pcpo *)
-
-val pcpoS = @{sort pcpo};
-fun mk_TFree s = TFree ("'" ^ s, pcpoS);
-
-
-(* basic types *)
-
-fun mk_btyp t (S, T) = Type (t, [S, T]);
-
-val cfun_arrow = @{type_name "cfun"};
-val op ->> = mk_btyp cfun_arrow;
-val mk_ssumT = mk_btyp (@{type_name "ssum"});
-val mk_sprodT = mk_btyp (@{type_name "sprod"});
-fun mk_uT T = Type (@{type_name u}, [T]);
-val trT = @{typ tr};
-val oneT = @{typ one};
-
-end;
--- a/src/Pure/General/sha1_polyml.ML Thu Mar 25 17:55:55 2010 +0100
+++ b/src/Pure/General/sha1_polyml.ML Thu Mar 25 17:56:31 2010 +0100
@@ -16,7 +16,7 @@
in (op ^) (pairself hex_digit (Integer.div_mod (Char.ord c) 16)) end
val lib_path =
- ("$ML_HOME/" ^ (if String.isSuffix "cygwin" ml_system then "sha1.dll" else "libsha1.so"))
+ ("$ML_HOME/" ^ (if String.isSuffix "cygwin" ml_platform then "sha1.dll" else "libsha1.so"))
|> Path.explode;
fun digest_external str =
@@ -24,7 +24,7 @@
val digest = CInterface.alloc 20 CInterface.Cchar;
val _ =
CInterface.call3 (CInterface.get_sym (File.platform_path lib_path) "sha1_buffer")
- (CInterface.STRING, CInterface.INT, CInterface.POINTER)
+ (CInterface.STRING, CInterface.LONG, CInterface.POINTER)
CInterface.POINTER (str, size str, CInterface.address digest);
in fold (suffix o hex_string digest) (0 upto 19) "" end;
--- a/src/Pure/Isar/isar_cmd.ML Thu Mar 25 17:55:55 2010 +0100
+++ b/src/Pure/Isar/isar_cmd.ML Thu Mar 25 17:56:31 2010 +0100
@@ -166,7 +166,7 @@
(* old-style axioms *)
-val add_axioms = fold (fn (b, ax) => snd o Specification.axiomatization_cmd [] [(b, [ax])]);
+val add_axioms = fold (snd oo Specification.axiom_cmd);
fun add_defs ((unchecked, overloaded), args) thy =
thy |> (if unchecked then PureThy.add_defs_unchecked_cmd else PureThy.add_defs_cmd) overloaded
--- a/src/Pure/Isar/specification.ML Thu Mar 25 17:55:55 2010 +0100
+++ b/src/Pure/Isar/specification.ML Thu Mar 25 17:56:31 2010 +0100
@@ -32,6 +32,8 @@
val axiomatization_cmd: (binding * string option * mixfix) list ->
(Attrib.binding * string list) list -> theory ->
(term list * thm list list) * theory
+ val axiom: Attrib.binding * term -> theory -> thm * theory
+ val axiom_cmd: Attrib.binding * string -> theory -> thm * theory
val definition:
(binding * typ option * mixfix) option * (Attrib.binding * term) ->
local_theory -> (term * (string * thm)) * local_theory
@@ -187,6 +189,9 @@
val axiomatization = gen_axioms false check_specification;
val axiomatization_cmd = gen_axioms true read_specification;
+fun axiom (b, ax) = axiomatization [] [(b, [ax])] #>> (hd o hd o snd);
+fun axiom_cmd (b, ax) = axiomatization_cmd [] [(b, [ax])] #>> (hd o hd o snd);
+
(* definition *)
--- a/src/Pure/ROOT.ML Thu Mar 25 17:55:55 2010 +0100
+++ b/src/Pure/ROOT.ML Thu Mar 25 17:56:31 2010 +0100
@@ -92,14 +92,14 @@
use "term_ord.ML";
use "term_subst.ML";
use "old_term.ML";
-use "logic.ML";
use "General/pretty.ML";
use "context.ML";
use "context_position.ML";
+use "sorts.ML";
+use "type.ML";
+use "logic.ML";
use "Syntax/lexicon.ML";
use "Syntax/simple_syntax.ML";
-use "sorts.ML";
-use "type.ML";
use "config.ML";
--- a/src/Pure/axclass.ML Thu Mar 25 17:55:55 2010 +0100
+++ b/src/Pure/axclass.ML Thu Mar 25 17:56:31 2010 +0100
@@ -514,6 +514,12 @@
local
+(* old-style axioms *)
+
+fun add_axiom (b, prop) =
+ Thm.add_axiom (b, prop) #->
+ (fn thm => PureThy.add_thm ((b, Drule.export_without_context thm), []));
+
fun axiomatize prep mk name add raw_args thy =
let
val args = prep thy raw_args;
@@ -521,7 +527,7 @@
val names = name args;
in
thy
- |> fold_map Drule.add_axiom (map Binding.name names ~~ specs)
+ |> fold_map add_axiom (map Binding.name names ~~ specs)
|-> fold add
end;
--- a/src/Pure/drule.ML Thu Mar 25 17:55:55 2010 +0100
+++ b/src/Pure/drule.ML Thu Mar 25 17:56:31 2010 +0100
@@ -77,7 +77,6 @@
val flexflex_unique: thm -> thm
val export_without_context: thm -> thm
val export_without_context_open: thm -> thm
- val add_axiom: (binding * term) -> theory -> thm * theory
val store_thm: binding -> thm -> thm
val store_standard_thm: binding -> thm -> thm
val store_thm_open: binding -> thm -> thm
@@ -321,12 +320,6 @@
#> Thm.close_derivation;
-(* old-style axioms *)
-
-fun add_axiom (b, prop) =
- Thm.add_axiom (b, prop) #-> (fn thm => PureThy.add_thm ((b, export_without_context thm), []));
-
-
(*Convert all Vars in a theorem to Frees. Also return a function for
reversing that operation. DOES NOT WORK FOR TYPE VARIABLES.
Similar code in type/freeze_thaw*)
--- a/src/Tools/Cache_IO/cache_io.ML Thu Mar 25 17:55:55 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,121 +0,0 @@
-(* Title: Tools/Cache_IO/cache_io.ML
- Author: Sascha Boehme, TU Muenchen
-
-Cache for output of external processes.
-*)
-
-signature CACHE_IO =
-sig
- val with_tmp_file: string -> (Path.T -> 'a) -> 'a
- val run: (Path.T -> string) -> Path.T -> string list
- val run': (Path.T -> Path.T -> string) -> Path.T -> string list * string list
-
- type cache
- val make: Path.T -> cache
- val cache_path_of: cache -> Path.T
- val cached: cache -> (Path.T -> string) -> Path.T -> string list
- val cached': cache -> (Path.T -> Path.T -> string) -> Path.T ->
- string list * string list
-end
-
-structure Cache_IO : CACHE_IO =
-struct
-
-fun with_tmp_file name f =
- let
- val path = File.tmp_path (Path.explode (name ^ serial_string ()))
- val x = Exn.capture f path
- val _ = try File.rm path
- in Exn.release x end
-
-fun run' make_cmd in_path =
- with_tmp_file "cache-io-" (fn out_path =>
- let
- val (out2, _) = bash_output (make_cmd in_path out_path)
- val out1 = the_default [] (try (rev o File.fold_lines cons out_path) [])
- in (out1, split_lines out2) end)
-
-fun run make_cmd = snd o run' (fn in_path => fn _ => make_cmd in_path)
-
-
-
-abstype cache = Cache of {
- path: Path.T,
- table: (int * (int * int * int) Symtab.table) Synchronized.var }
-with
-
-
-fun cache_path_of (Cache {path, ...}) = path
-
-
-fun load cache_path =
- let
- fun err () = error ("Cache IO: corrupted cache file: " ^
- File.shell_path cache_path)
-
- fun int_of_string s =
- (case read_int (explode s) of
- (i, []) => i
- | _ => err ())
-
- fun split line =
- (case space_explode " " line of
- [key, len1, len2] => (key, int_of_string len1, int_of_string len2)
- | _ => err ())
-
- fun parse line ((i, l), tab) =
- if i = l
- then
- let val (key, l1, l2) = split line
- in ((i+1, l+l1+l2+1), Symtab.update (key, (i+1, l1, l2)) tab) end
- else ((i+1, l), tab)
- in apfst fst (File.fold_lines parse cache_path ((1, 1), Symtab.empty)) end
-
-fun make path =
- let val table = if File.exists path then load path else (1, Symtab.empty)
- in Cache {path=path, table=Synchronized.var (Path.implode path) table} end
-
-
-fun get_hash_key path =
- let
- val arg = File.shell_path path
- val (out, res) = bash_output (getenv "COMPUTE_HASH_KEY" ^ " " ^ arg)
- in
- if res = 0 then hd (split_lines out)
- else error ("Cache IO: failed to generate hash key for file " ^ arg)
- end
-
-fun load_cached_result cache_path (p, len1, len2) =
- let
- fun load line (i, xsp) =
- if i < p then (i+1, xsp)
- else if i < p + len1 then (i+1, apfst (cons line) xsp)
- else if i < p + len2 then (i+1, apsnd (cons line) xsp)
- else (i, xsp)
- in pairself rev (snd (File.fold_lines load cache_path (1, ([], [])))) end
-
-fun cached' (Cache {path=cache_path, table}) make_cmd in_path =
- let val key = get_hash_key in_path
- in
- (case Symtab.lookup (snd (Synchronized.value table)) key of
- SOME pos => load_cached_result cache_path pos
- | NONE =>
- let
- val res as (out, err) = run' make_cmd in_path
- val (l1, l2) = pairself length res
- val header = key ^ " " ^ string_of_int l1 ^ " " ^ string_of_int l2
- val lines = map (suffix "\n") (header :: out @ err)
-
- val _ = Synchronized.change table (fn (p, tab) =>
- if Symtab.defined tab key then (p, tab)
- else
- let val _ = File.append_list cache_path lines
- in (p+l1+l2+1, Symtab.update (key, (p+1, l1, l2)) tab) end)
- in res end)
- end
-
-fun cached cache make_cmd =
- snd o cached' cache (fn in_path => fn _ => make_cmd in_path)
-
-end
-end
--- a/src/Tools/Cache_IO/etc/settings Thu Mar 25 17:55:55 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-ISABELLE_CACHE_IO="$COMPONENT"
-
-COMPUTE_HASH_KEY="$ISABELLE_CACHE_IO/lib/scripts/compute_hash_key"
-
--- a/src/Tools/Cache_IO/lib/scripts/compute_hash_key Thu Mar 25 17:55:55 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,24 +0,0 @@
-#!/usr/bin/env perl
-#
-# Author: Sascha Boehme, TU Muenchen
-#
-# Compute MD5 hash key.
-
-use strict;
-use warnings;
-use Digest::MD5;
-
-
-# argument
-
-my $file = $ARGV[0];
-
-
-# compute MD5 hash key
-
-my $md5 = Digest::MD5->new;
-open FILE, "<$file" or die "ERROR: Failed to open '$file' ($!)";
-$md5->addfile(*FILE);
-close FILE;
-print $md5->b64digest . "\n";
-
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/cache_io.ML Thu Mar 25 17:56:31 2010 +0100
@@ -0,0 +1,112 @@
+(* Title: Tools/cache_io.ML
+ Author: Sascha Boehme, TU Muenchen
+
+Cache for output of external processes.
+*)
+
+signature CACHE_IO =
+sig
+ val with_tmp_file: string -> (Path.T -> 'a) -> 'a
+ val run: (Path.T -> string) -> Path.T -> string list
+ val run': (Path.T -> Path.T -> string) -> Path.T -> string list * string list
+
+ type cache
+ val make: Path.T -> cache
+ val cache_path_of: cache -> Path.T
+ val cached: cache -> (Path.T -> string) -> Path.T -> string list
+ val cached': cache -> (Path.T -> Path.T -> string) -> Path.T ->
+ string list * string list
+end
+
+structure Cache_IO : CACHE_IO =
+struct
+
+fun with_tmp_file name f =
+ let
+ val path = File.tmp_path (Path.explode (name ^ serial_string ()))
+ val x = Exn.capture f path
+ val _ = try File.rm path
+ in Exn.release x end
+
+fun run' make_cmd in_path =
+ with_tmp_file "cache-io-" (fn out_path =>
+ let
+ val (out2, _) = bash_output (make_cmd in_path out_path)
+ val out1 = the_default [] (try (rev o File.fold_lines cons out_path) [])
+ in (out1, split_lines out2) end)
+
+fun run make_cmd = snd o run' (fn in_path => fn _ => make_cmd in_path)
+
+
+
+abstype cache = Cache of {
+ path: Path.T,
+ table: (int * (int * int * int) Symtab.table) Synchronized.var }
+with
+
+
+fun cache_path_of (Cache {path, ...}) = path
+
+
+fun load cache_path =
+ let
+ fun err () = error ("Cache IO: corrupted cache file: " ^
+ File.shell_path cache_path)
+
+ fun int_of_string s =
+ (case read_int (explode s) of
+ (i, []) => i
+ | _ => err ())
+
+ fun split line =
+ (case space_explode " " line of
+ [key, len1, len2] => (key, int_of_string len1, int_of_string len2)
+ | _ => err ())
+
+ fun parse line ((i, l), tab) =
+ if i = l
+ then
+ let val (key, l1, l2) = split line
+ in ((i+1, l+l1+l2+1), Symtab.update (key, (i+1, l1, l2)) tab) end
+ else ((i+1, l), tab)
+ in apfst fst (File.fold_lines parse cache_path ((1, 1), Symtab.empty)) end
+
+fun make path =
+ let val table = if File.exists path then load path else (1, Symtab.empty)
+ in Cache {path=path, table=Synchronized.var (Path.implode path) table} end
+
+
+fun load_cached_result cache_path (p, len1, len2) =
+ let
+ fun load line (i, xsp) =
+ if i < p then (i+1, xsp)
+ else if i < p + len1 then (i+1, apfst (cons line) xsp)
+ else if i < p + len2 then (i+1, apsnd (cons line) xsp)
+ else (i, xsp)
+ in pairself rev (snd (File.fold_lines load cache_path (1, ([], [])))) end
+
+fun cached' (Cache {path=cache_path, table}) make_cmd in_path =
+ let val key = SHA1.digest (File.read in_path)
+ in
+ (case Symtab.lookup (snd (Synchronized.value table)) key of
+ SOME pos => load_cached_result cache_path pos
+ | NONE =>
+ let
+ val res as (out, err) = run' make_cmd in_path
+ val (l1, l2) = pairself length res
+ val header = key ^ " " ^ string_of_int l1 ^ " " ^ string_of_int l2
+ val lines = map (suffix "\n") (header :: out @ err)
+
+ val _ = Synchronized.change table (fn (p, tab) =>
+ if Symtab.defined tab key then (p, tab)
+ else
+ let val _ = File.append_list cache_path lines
+ in (p+l1+l2+1, Symtab.update (key, (p+1, l1, l2)) tab) end)
+ in res end)
+ end
+
+fun cached cache make_cmd =
+ snd o cached' cache (fn in_path => fn _ => make_cmd in_path)
+
+end
+end