--- a/CONTRIBUTORS Mon Mar 30 10:58:08 2015 +0200
+++ b/CONTRIBUTORS Mon Mar 30 11:36:30 2015 +0200
@@ -6,6 +6,9 @@
Contributions to this Isabelle version
--------------------------------------
+* March 2015: Jasmin Blanchette, Inria & LORIA & MPII, Mathias Fleury, MPII, and Dmitriy Traytel, TUM
+ More multiset theorems, syntax, and operations.
+
* December 2014: Johannes Hölzl, Manuel Eberl, Sudeep Kanav, TUM and Jeremy Avigad, Luke Serafin, CMU
Various integration theorems: mostly integration on intervals and substitution.
--- a/NEWS Mon Mar 30 10:58:08 2015 +0200
+++ b/NEWS Mon Mar 30 11:36:30 2015 +0200
@@ -61,15 +61,19 @@
*** Pure ***
-* Proof methods with explicit instantiation ("rule_tac", "subgoal_tac"
-etc.) allow an optional context of local variables ('for' declaration):
-these variables become schematic in the instantiated theorem. This
-behaviour is analogous to 'for' in attributes "where" and "of".
-
* Explicit instantiation via attributes "where", "of", and proof methods
"rule_tac" with derivatives like "subgoal_tac" etc. admit dummy patterns
("_") that stand for anonymous local variables.
+* Proof methods with explicit instantiation ("rule_tac", "subgoal_tac"
+etc.) allow an optional context of local variables ('for' declaration):
+these variables become schematic in the instantiated theorem; this
+behaviour is analogous to 'for' in attributes "where" and "of".
+Configuration option rule_insts_schematic (default false) controls use
+of schematic variables outside the context. Minor INCOMPATIBILITY,
+declare rule_insts_schematic = true temporarily and update to use local
+variable declarations or dummy patterns instead.
+
* Command "class_deps" takes optional sort arguments constraining
the search space.
@@ -85,6 +89,9 @@
*** HOL ***
+* Type classes cancel_ab_semigroup_add / cancel_monoid_add specify
+explicit additive inverse operation. INCOMPATIBILITY.
+
* New proof method "rewrite" (in ~~/src/HOL/Library/Rewrite) for
single-step rewriting with subterm selection based on patterns.
@@ -257,6 +264,24 @@
The "sos_cert" functionality is invoked as "sos" with additional
argument. Minor INCOMPATIBILITY.
+* Theory "Library/Multiset":
+ - Introduced "replicate_mset" operation.
+ - Introduced alternative characterizations of the multiset ordering in
+ "Library/Multiset_Order".
+ - Renamed
+ in_multiset_of ~> in_multiset_in_set
+ INCOMPATIBILITY.
+ - Added attributes:
+ image_mset.id [simp]
+ image_mset_id [simp]
+ elem_multiset_of_set [simp, intro]
+ comp_fun_commute_plus_mset [simp]
+ comp_fun_commute.fold_mset_insert [OF comp_fun_commute_plus_mset, simp]
+ in_mset_fold_plus_iff [iff]
+ set_of_Union_mset [simp]
+ in_Union_mset_iff [iff]
+ INCOMPATIBILITY.
+
* HOL-Decision_Procs:
- New counterexample generator quickcheck[approximation] for
inequalities of transcendental functions.
--- a/etc/components Mon Mar 30 10:58:08 2015 +0200
+++ b/etc/components Mon Mar 30 11:36:30 2015 +0200
@@ -4,6 +4,7 @@
src/HOL/Mirabelle
src/HOL/Mutabelle
src/HOL/Library/Sum_of_Squares
+src/HOL/SPARK
src/HOL/Tools
src/HOL/Tools/ATP
src/HOL/TPTP
--- a/src/CCL/ex/Stream.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/CCL/ex/Stream.thy Mon Mar 30 11:36:30 2015 +0200
@@ -120,7 +120,7 @@
lemma iter2Blemma:
"n:Nat \<Longrightarrow>
map(f) ^ n ` iter2(f,a) = (f ^ n ` a) $ (map(f) ^ n ` map(f,iter2(f,a)))"
- apply (rule_tac P = "\<lambda>x. ?lhs (x) = ?rhs" in iter2B [THEN ssubst])
+ apply (rule_tac P = "\<lambda>x. lhs(x) = rhs" for lhs rhs in iter2B [THEN ssubst])
apply (simp add: nmapBcons)
done
--- a/src/Doc/Datatypes/Datatypes.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Doc/Datatypes/Datatypes.thy Mon Mar 30 11:36:30 2015 +0200
@@ -865,25 +865,25 @@
\item[@{text "t."}\hthm{case_transfer} @{text "[transfer_rule]"}\rm:] ~ \\
@{thm list.case_transfer[no_vars]} \\
The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin
-(Section~\ref{ssec:transfer}).
+(Section~\ref{ssec:transfer}) for type constructors with no dead type arguments.
\item[@{text "t."}\hthm{sel_transfer} @{text "[transfer_rule]"}\rm:] ~ \\
This property is missing for @{typ "'a list"} because there is no common
selector to all constructors. \\
The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin
-(Section~\ref{ssec:transfer}).
+(Section~\ref{ssec:transfer}) for type constructors with no dead type arguments.
\item[@{text "t."}\hthm{ctr_transfer} @{text "[transfer_rule]"}\rm:] ~ \\
@{thm list.ctr_transfer(1)[no_vars]} \\
@{thm list.ctr_transfer(2)[no_vars]} \\
The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin
-(Section~\ref{ssec:transfer}).
+(Section~\ref{ssec:transfer}) for type constructors with no dead type arguments.
\item[@{text "t."}\hthm{disc_transfer} @{text "[transfer_rule]"}\rm:] ~ \\
@{thm list.disc_transfer(1)[no_vars]} \\
@{thm list.disc_transfer(2)[no_vars]} \\
The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin
-(Section~\ref{ssec:transfer}).
+(Section~\ref{ssec:transfer}) for type constructors with no dead type arguments.
\item[@{text "t."}\hthm{set} @{text "[simp, code]"}\rm:] ~ \\
@{thm list.set(1)[no_vars]} \\
@@ -958,7 +958,7 @@
\item[@{text "t."}\hthm{set_transfer} @{text "[transfer_rule]"}\rm:] ~ \\
@{thm list.set_transfer[no_vars]} \\
The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin
-(Section~\ref{ssec:transfer}).
+(Section~\ref{ssec:transfer}) for type constructors with no dead type arguments.
\item[@{text "t."}\hthm{map_cong0}\rm:] ~ \\
@{thm list.map_cong0[no_vars]}
@@ -981,7 +981,7 @@
\item[@{text "t."}\hthm{map_transfer} @{text "[transfer_rule]"}\rm:] ~ \\
@{thm list.map_transfer[no_vars]} \\
The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin
-(Section~\ref{ssec:transfer}).
+(Section~\ref{ssec:transfer}) for type constructors with no dead type arguments.
\item[@{text "t."}\hthm{rel_compp} @{text "[relator_distr]"}\rm:] ~ \\
@{thm list.rel_compp[no_vars]} \\
@@ -1012,7 +1012,7 @@
\item[@{text "t."}\hthm{rel_transfer} @{text "[transfer_rule]"}\rm:] ~ \\
@{thm list.rel_transfer[no_vars]} \\
The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin
-(Section~\ref{ssec:transfer}).
+(Section~\ref{ssec:transfer}) for type constructors with no dead type arguments.
\end{description}
\end{indentblock}
@@ -1053,7 +1053,7 @@
\item[@{text "t."}\hthm{rec_transfer} @{text "[transfer_rule]"}\rm:] ~ \\
@{thm list.rec_transfer[no_vars]} \\
The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin
-(Section~\ref{ssec:transfer}).
+(Section~\ref{ssec:transfer}) for type constructors with no dead type arguments.
\end{description}
\end{indentblock}
@@ -1928,7 +1928,7 @@
\item[@{text "t."}\hthm{corec_transfer} @{text "[transfer_rule]"}\rm:] ~ \\
@{thm llist.corec_transfer[no_vars]} \\
The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin
-(Section~\ref{ssec:transfer}).
+(Section~\ref{ssec:transfer}) for type constructors with no dead type arguments.
\end{description}
\end{indentblock}
@@ -3066,7 +3066,8 @@
the \hthm{transfer} plugin generates a predicator @{text "t.pred_t"} and
properties that guide the Transfer tool.
-The plugin derives the following properties:
+For types with no dead type arguments (and at least one live type argument), the
+plugin derives the following properties:
\begin{indentblock}
\begin{description}
--- a/src/Doc/Isar_Ref/HOL_Specific.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Doc/Isar_Ref/HOL_Specific.thy Mon Mar 30 11:36:30 2015 +0200
@@ -2338,7 +2338,7 @@
;
@@{method (HOL) induct_tac} @{syntax goal_spec}? (@{syntax insts} * @'and') rule?
;
- @@{method (HOL) ind_cases} (@{syntax prop}+) (@'for' (@{syntax name}+))?
+ @@{method (HOL) ind_cases} (@{syntax prop}+) @{syntax for_fixes}
;
@@{command (HOL) inductive_cases} (@{syntax thmdecl}? (@{syntax prop}+) + @'and')
;
--- a/src/Doc/antiquote_setup.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Doc/antiquote_setup.ML Mon Mar 30 11:36:30 2015 +0200
@@ -74,7 +74,7 @@
(Input.source_content source, ML_Lex.read_source false source);
fun index_ml name kind ml = Thy_Output.antiquotation name
- (Scan.lift (Args.text_source_position -- Scan.option (Args.colon |-- Args.text_source_position)))
+ (Scan.lift (Args.text_input -- Scan.option (Args.colon |-- Args.text_input)))
(fn {context = ctxt, ...} => fn (source1, opt_source2) =>
let
val (txt1, toks1) = prep_ml source1;
--- a/src/HOL/Algebra/Sylow.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Algebra/Sylow.thy Mon Mar 30 11:36:30 2015 +0200
@@ -182,7 +182,7 @@
apply (erule H_into_carrier_G)
apply (rule H_not_empty)
apply (simp add: H_def, clarify)
-apply (erule_tac P = "%z. ?lhs(z) = M1" in subst)
+apply (erule_tac P = "%z. lhs(z) = M1" for lhs in subst)
apply (simp add: coset_mult_assoc )
apply (blast intro: H_m_closed)
done
--- a/src/HOL/Auth/KerberosIV_Gets.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Auth/KerberosIV_Gets.thy Mon Mar 30 11:36:30 2015 +0200
@@ -1238,11 +1238,12 @@
txt{*K5. Not clear how this step could be integrated with the main
simplification step. Done in KerberosV.thy*}
apply clarify
-apply (erule_tac V = "Says Aa Tgs ?X \<in> set ?evs" in thin_rl)
+apply (erule_tac V = "Says Aa Tgs X \<in> set evs" for X evs in thin_rl)
apply (frule Gets_imp_knows_Spy [THEN parts.Inj, THEN servK_notin_authKeysD])
apply (assumption, assumption, blast, assumption)
apply (simp add: analz_insert_freshK2)
-by (blast dest: Says_imp_spies [THEN parts.Inj] Key_unique_SesKey intro: less_SucI)
+apply (blast dest: Key_unique_SesKey intro: less_SucI)
+done
text{* In the real world Tgs can't check wheter authK is secure! *}
--- a/src/HOL/Auth/Yahalom.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Auth/Yahalom.thy Mon Mar 30 11:36:30 2015 +0200
@@ -403,7 +403,7 @@
txt{*YM3*}
apply blast
txt{*YM4*}
-apply (erule_tac V = "\<forall>KK. ?P KK" in thin_rl, clarify)
+apply (erule_tac V = "\<forall>KK. P KK" for P in thin_rl, clarify)
txt{*If @{prop "A \<in> bad"} then @{term NBa} is known, therefore
@{prop "NBa \<noteq> NB"}. Previous two steps make the next step
faster.*}
--- a/src/HOL/Bali/AxCompl.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Bali/AxCompl.thy Mon Mar 30 11:36:30 2015 +0200
@@ -71,7 +71,7 @@
apply (subgoal_tac
"nyinitcls G (x,s) = insert C (nyinitcls G (x,init_class_obj G C s))")
apply clarsimp
-apply (erule_tac V="nyinitcls G (x, s) = ?rhs" in thin_rl)
+apply (erule_tac V="nyinitcls G (x, s) = rhs" for rhs in thin_rl)
apply (rule card_Suc_lemma [OF _ _ finite_nyinitcls])
apply (auto dest!: not_initedD elim!:
simp add: nyinitcls_def inited_def split add: split_if_asm)
--- a/src/HOL/Bali/AxExample.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Bali/AxExample.thy Mon Mar 30 11:36:30 2015 +0200
@@ -71,13 +71,13 @@
apply (rule_tac P' = "Normal (\<lambda>Y s Z. arr_inv (snd s))" in conseq1)
prefer 2
apply clarsimp
-apply (rule_tac Q' = "(\<lambda>Y s Z. ?Q Y s Z)\<leftarrow>=False\<down>=\<diamondsuit>" in conseq2)
+apply (rule_tac Q' = "(\<lambda>Y s Z. Q Y s Z)\<leftarrow>=False\<down>=\<diamondsuit>" and Q = Q for Q in conseq2)
prefer 2
apply simp
apply (tactic "ax_tac @{context} 1" (* While *))
prefer 2
apply (rule ax_impossible [THEN conseq1], clarsimp)
-apply (rule_tac P' = "Normal ?P" in conseq1)
+apply (rule_tac P' = "Normal P" and P = P for P in conseq1)
prefer 2
apply clarsimp
apply (tactic "ax_tac @{context} 1")
@@ -106,7 +106,7 @@
apply (unfold DynT_prop_def)
apply (simp (no_asm))
apply (intro strip)
-apply (rule_tac P' = "Normal ?P" in conseq1)
+apply (rule_tac P' = "Normal P" and P = P for P in conseq1)
apply (tactic "ax_tac @{context} 1" (* Methd *))
apply (rule ax_thin [OF _ empty_subsetI])
apply (simp (no_asm) add: body_def2)
@@ -185,7 +185,7 @@
apply (simp (no_asm))
apply (unfold arr_viewed_from_def)
apply (rule allI)
-apply (rule_tac P' = "Normal ?P" in conseq1)
+apply (rule_tac P' = "Normal P" and P = P for P in conseq1)
apply (tactic {* simp_tac (@{context} delloop "split_all_tac") 1 *})
apply (tactic "ax_tac @{context} 1")
apply (tactic "ax_tac @{context} 1")
@@ -268,7 +268,7 @@
apply (rule ax_subst_Var_allI)
apply (tactic {* inst1_tac @{context} "P'" "\<lambda>b Y ba Z vf. \<lambda>Y (x,s) Z. x=None \<and> snd vf = snd (lvar i s)" [] *})
apply (rule allI)
-apply (rule_tac P' = "Normal ?P" in conseq1)
+apply (rule_tac P' = "Normal P" and P = P for P in conseq1)
prefer 2
apply clarsimp
apply (tactic "ax_tac @{context} 1")
--- a/src/HOL/Bali/AxSem.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Bali/AxSem.thy Mon Mar 30 11:36:30 2015 +0200
@@ -1055,7 +1055,7 @@
apply safe
apply (erule spec)
apply (rule ax_escape, clarsimp)
-apply (erule_tac V = "?P \<longrightarrow> ?Q" in thin_rl)
+apply (erule_tac V = "P \<longrightarrow> Q" for P Q in thin_rl)
apply (drule spec,drule spec,drule spec, erule conseq12)
apply (force simp add: init_lvars_def Let_def)
done
--- a/src/HOL/Bali/Example.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Bali/Example.thy Mon Mar 30 11:36:30 2015 +0200
@@ -1259,18 +1259,18 @@
apply (simp (no_asm_use))
apply (drule (1) alloc_one,clarsimp)
apply (rule eval_Is (* ;; *))
-apply (erule_tac V = "the (new_Addr ?h) = c" in thin_rl)
-apply (erule_tac [2] V = "new_Addr ?h = Some a" in thin_rl)
-apply (erule_tac [2] V = "atleast_free ?h four" in thin_rl)
+apply (erule_tac V = "the (new_Addr _) = c" in thin_rl)
+apply (erule_tac [2] V = "new_Addr _ = Some a" in thin_rl)
+apply (erule_tac [2] V = "atleast_free _ four" in thin_rl)
apply (rule eval_Is (* Expr *))
apply (rule eval_Is (* Ass *))
apply (rule eval_Is (* LVar *))
apply (rule eval_Is (* NewC *))
(* begin init Ext *)
-apply (erule_tac V = "the (new_Addr ?h) = b" in thin_rl)
-apply (erule_tac V = "atleast_free ?h three" in thin_rl)
-apply (erule_tac [2] V = "atleast_free ?h four" in thin_rl)
-apply (erule_tac [2] V = "new_Addr ?h = Some a" in thin_rl)
+apply (erule_tac V = "the (new_Addr _) = b" in thin_rl)
+apply (erule_tac V = "atleast_free _ three" in thin_rl)
+apply (erule_tac [2] V = "atleast_free _ four" in thin_rl)
+apply (erule_tac [2] V = "new_Addr _ = Some a" in thin_rl)
apply (rule eval_Is (* Init Ext *))
apply (simp)
apply (rule conjI)
@@ -1309,7 +1309,7 @@
apply (drule alloc_one)
apply (simp (no_asm_simp))
apply clarsimp
-apply (erule_tac V = "atleast_free ?h three" in thin_rl)
+apply (erule_tac V = "atleast_free _ three" in thin_rl)
apply (drule_tac x = "a" in new_AddrD2 [THEN spec])
apply (simp (no_asm_use))
apply (rule eval_Is (* Try *))
@@ -1361,7 +1361,7 @@
apply (simp add: gupd_def lupd_def obj_ty_def split del: split_if)
apply (drule alloc_one [THEN conjunct1])
apply (simp (no_asm_simp))
-apply (erule_tac V = "atleast_free ?h two" in thin_rl)
+apply (erule_tac V = "atleast_free _ two" in thin_rl)
apply (drule_tac x = "a" in new_AddrD2 [THEN spec])
apply simp
apply (rule eval_Is (* While *))
--- a/src/HOL/Bali/State.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Bali/State.thy Mon Mar 30 11:36:30 2015 +0200
@@ -599,13 +599,13 @@
where "store == snd"
lemma single_stateE: "\<forall>Z. Z = (s::state) \<Longrightarrow> False"
-apply (erule_tac x = "(Some k,y)" in all_dupE)
-apply (erule_tac x = "(None,y)" in allE)
+apply (erule_tac x = "(Some k,y)" for k y in all_dupE)
+apply (erule_tac x = "(None,y)" for y in allE)
apply clarify
done
lemma state_not_single: "All (op = (x::state)) \<Longrightarrow> R"
-apply (drule_tac x = "(if abrupt x = None then Some ?x else None,?y)" in spec)
+apply (drule_tac x = "(if abrupt x = None then Some x' else None, y)" for x' y in spec)
apply clarsimp
done
@@ -823,4 +823,3 @@
end
-
--- a/src/HOL/Bali/WellType.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Bali/WellType.thy Mon Mar 30 11:36:30 2015 +0200
@@ -660,14 +660,14 @@
(* 17 subgoals *)
apply (tactic {* ALLGOALS (fn i =>
if i = 11 then EVERY'
- [Rule_Insts.thin_tac @{context} "?E,dt\<Turnstile>e0\<Colon>-PrimT Boolean" [],
- Rule_Insts.thin_tac @{context} "?E,dt\<Turnstile>e1\<Colon>-?T1" [],
- Rule_Insts.thin_tac @{context} "?E,dt\<Turnstile>e2\<Colon>-?T2" []] i
- else Rule_Insts.thin_tac @{context} "All ?P" [] i) *})
+ [Rule_Insts.thin_tac @{context} "E,dt\<Turnstile>e0\<Colon>-PrimT Boolean" [(@{binding E}, NONE, NoSyn)],
+ Rule_Insts.thin_tac @{context} "E,dt\<Turnstile>e1\<Colon>-T1" [(@{binding E}, NONE, NoSyn), (@{binding T1}, NONE, NoSyn)],
+ Rule_Insts.thin_tac @{context} "E,dt\<Turnstile>e2\<Colon>-T2" [(@{binding E}, NONE, NoSyn), (@{binding T2}, NONE, NoSyn)]] i
+ else Rule_Insts.thin_tac @{context} "All P" [(@{binding P}, NONE, NoSyn)] i) *})
(*apply (safe del: disjE elim!: wt_elim_cases)*)
apply (tactic {*ALLGOALS (eresolve_tac @{context} @{thms wt_elim_cases})*})
apply (simp_all (no_asm_use) split del: split_if_asm)
-apply (erule_tac [12] V = "All ?P" in thin_rl) (* Call *)
+apply (erule_tac [12] V = "All P" for P in thin_rl) (* Call *)
apply (blast del: equalityCE dest: sym [THEN trans])+
done
--- a/src/HOL/Basic_BNF_LFPs.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Basic_BNF_LFPs.thy Mon Mar 30 11:36:30 2015 +0200
@@ -99,8 +99,7 @@
declare prod.size[no_atp]
-lemma size_nat[simp, code]: "size (n\<Colon>nat) = n"
- by (induct n) simp_all
+lemmas size_nat = size_nat_def
hide_const (open) xtor ctor_rec
--- a/src/HOL/Code_Numeral.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Code_Numeral.thy Mon Mar 30 11:36:30 2015 +0200
@@ -217,9 +217,7 @@
instance integer :: semiring_numeral_div
by intro_classes (transfer,
- fact semiring_numeral_div_class.diff_invert_add1
- semiring_numeral_div_class.le_add_diff_inverse2
- semiring_numeral_div_class.mult_div_cancel
+ fact semiring_numeral_div_class.le_add_diff_inverse2
semiring_numeral_div_class.div_less
semiring_numeral_div_class.mod_less
semiring_numeral_div_class.div_positive
--- a/src/HOL/Codegenerator_Test/Candidates.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Codegenerator_Test/Candidates.thy Mon Mar 30 11:36:30 2015 +0200
@@ -13,11 +13,12 @@
begin
setup \<open>
+fn thy =>
let
- val tycos = (#log_types o Type.rep_tsig o Sign.tsig_of) @{theory};
- val consts = map_filter (try (curry (Axclass.param_of_inst @{theory})
+ val tycos = Sign.logical_types thy;
+ val consts = map_filter (try (curry (Axclass.param_of_inst thy)
@{const_name "Quickcheck_Narrowing.partial_term_of"})) tycos;
-in fold Code.del_eqns consts end
+in fold Code.del_eqns consts thy end
\<close> -- \<open>drop technical stuff from @{text Quickcheck_Narrowing} which is tailored towards Haskell\<close>
inductive sublist :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
--- a/src/HOL/Codegenerator_Test/Generate_Efficient_Datastructures.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Codegenerator_Test/Generate_Efficient_Datastructures.thy Mon Mar 30 11:36:30 2015 +0200
@@ -12,11 +12,12 @@
begin
setup \<open>
+fn thy =>
let
- val tycos = (#log_types o Type.rep_tsig o Sign.tsig_of) @{theory};
- val consts = map_filter (try (curry (Axclass.param_of_inst @{theory})
+ val tycos = Sign.logical_types thy;
+ val consts = map_filter (try (curry (Axclass.param_of_inst thy)
@{const_name "Quickcheck_Narrowing.partial_term_of"})) tycos;
-in fold Code.del_eqns consts end
+in fold Code.del_eqns consts thy end
\<close> -- \<open>drop technical stuff from @{text Quickcheck_Narrowing} which is tailored towards Haskell\<close>
(*
--- a/src/HOL/Decision_Procs/Approximation.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Decision_Procs/Approximation.thy Mon Mar 30 11:36:30 2015 +0200
@@ -3583,130 +3583,31 @@
in Thm.global_cterm_of thy (Logic.mk_equals (t, normalize t)) end
*}
-ML {*
- fun reorder_bounds_tac prems i =
- let
- fun variable_of_bound (Const (@{const_name Trueprop}, _) $
- (Const (@{const_name Set.member}, _) $
- Free (name, _) $ _)) = name
- | variable_of_bound (Const (@{const_name Trueprop}, _) $
- (Const (@{const_name HOL.eq}, _) $
- Free (name, _) $ _)) = name
- | variable_of_bound t = raise TERM ("variable_of_bound", [t])
-
- val variable_bounds
- = map (`(variable_of_bound o Thm.prop_of)) prems
-
- fun add_deps (name, bnds)
- = Graph.add_deps_acyclic (name,
- remove (op =) name (Term.add_free_names (Thm.prop_of bnds) []))
-
- val order = Graph.empty
- |> fold Graph.new_node variable_bounds
- |> fold add_deps variable_bounds
- |> Graph.strong_conn |> map the_single |> rev
- |> map_filter (AList.lookup (op =) variable_bounds)
-
- fun prepend_prem th tac
- = tac THEN rtac (th RSN (2, @{thm mp})) i
- in
- fold prepend_prem order all_tac
- end
-
- fun approximation_conv ctxt ct =
- approximation_oracle (Proof_Context.theory_of ctxt, Thm.term_of ct |> tap (tracing o Syntax.string_of_term ctxt));
-
- fun approximate ctxt t =
- approximation_oracle (Proof_Context.theory_of ctxt, t)
- |> Thm.prop_of |> Logic.dest_equals |> snd;
-
- (* Should be in HOL.thy ? *)
- fun gen_eval_tac conv ctxt = CONVERSION
- (Object_Logic.judgment_conv (Conv.params_conv (~1) (K (Conv.concl_conv (~1) conv)) ctxt))
- THEN' rtac TrueI
-
- val form_equations = @{thms interpret_form_equations};
-
- fun rewrite_interpret_form_tac ctxt prec splitting taylor i st = let
- fun lookup_splitting (Free (name, _))
- = case AList.lookup (op =) splitting name
- of SOME s => HOLogic.mk_number @{typ nat} s
- | NONE => @{term "0 :: nat"}
- val vs = nth (Thm.prems_of st) (i - 1)
- |> Logic.strip_imp_concl
- |> HOLogic.dest_Trueprop
- |> Term.strip_comb |> snd |> List.last
- |> HOLogic.dest_list
- val p = prec
- |> HOLogic.mk_number @{typ nat}
- |> Thm.cterm_of ctxt
- in case taylor
- of NONE => let
- val n = vs |> length
- |> HOLogic.mk_number @{typ nat}
- |> Thm.cterm_of ctxt
- val s = vs
- |> map lookup_splitting
- |> HOLogic.mk_list @{typ nat}
- |> Thm.cterm_of ctxt
- in
- (rtac (Thm.instantiate ([], [(@{cpat "?n::nat"}, n),
- (@{cpat "?prec::nat"}, p),
- (@{cpat "?ss::nat list"}, s)])
- @{thm "approx_form"}) i
- THEN simp_tac (put_simpset (simpset_of @{context}) ctxt) i) st
- end
-
- | SOME t =>
- if length vs <> 1
- then raise (TERM ("More than one variable used for taylor series expansion", [Thm.prop_of st]))
- else let
- val t = t
- |> HOLogic.mk_number @{typ nat}
- |> Thm.cterm_of ctxt
- val s = vs |> map lookup_splitting |> hd
- |> Thm.cterm_of ctxt
- in
- rtac (Thm.instantiate ([], [(@{cpat "?s::nat"}, s),
- (@{cpat "?t::nat"}, t),
- (@{cpat "?prec::nat"}, p)])
- @{thm "approx_tse_form"}) i st
- end
- end
-
- val free = Args.context -- Args.term >> (fn (_, Free (n, _)) => n | (ctxt, t) =>
- error ("Bad free variable: " ^ Syntax.string_of_term ctxt t));
-*}
-
lemma intervalE: "a \<le> x \<and> x \<le> b \<Longrightarrow> \<lbrakk> x \<in> { a .. b } \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by auto
lemma meta_eqE: "x \<equiv> a \<Longrightarrow> \<lbrakk> x = a \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by auto
+ML_file "approximation.ML"
+
method_setup approximation = {*
- Scan.lift Parse.nat
- --
- Scan.optional (Scan.lift (Args.$$$ "splitting" |-- Args.colon)
- |-- Parse.and_list' (free --| Scan.lift (Args.$$$ "=") -- Scan.lift Parse.nat)) []
- --
- Scan.option (Scan.lift (Args.$$$ "taylor" |-- Args.colon)
- |-- (free |-- Scan.lift (Args.$$$ "=") |-- Scan.lift Parse.nat))
- >>
- (fn ((prec, splitting), taylor) => fn ctxt =>
- SIMPLE_METHOD' (fn i =>
- REPEAT (FIRST' [etac @{thm intervalE},
- etac @{thm meta_eqE},
- rtac @{thm impI}] i)
- THEN Subgoal.FOCUS (fn {prems, ...} => reorder_bounds_tac prems i) ctxt i
- THEN DETERM (TRY (filter_prems_tac ctxt (K false) i))
- THEN DETERM (Reification.tac ctxt form_equations NONE i)
- THEN rewrite_interpret_form_tac ctxt prec splitting taylor i
- THEN gen_eval_tac (approximation_conv ctxt) ctxt i))
+ let val free = Args.context -- Args.term >> (fn (_, Free (n, _)) => n | (ctxt, t) =>
+ error ("Bad free variable: " ^ Syntax.string_of_term ctxt t));
+ in
+ Scan.lift Parse.nat
+ --
+ Scan.optional (Scan.lift (Args.$$$ "splitting" |-- Args.colon)
+ |-- Parse.and_list' (free --| Scan.lift (Args.$$$ "=") -- Scan.lift Parse.nat)) []
+ --
+ Scan.option (Scan.lift (Args.$$$ "taylor" |-- Args.colon)
+ |-- (free |-- Scan.lift (Args.$$$ "=") |-- Scan.lift Parse.nat))
+ >>
+ (fn ((prec, splitting), taylor) => fn ctxt =>
+ SIMPLE_METHOD' (Approximation.approximation_tac prec splitting taylor ctxt))
+ end
*} "real number approximation"
-ML_file "approximation.ML"
-
section "Quickcheck Generator"
--- a/src/HOL/Decision_Procs/Polynomial_List.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Decision_Procs/Polynomial_List.thy Mon Mar 30 11:36:30 2015 +0200
@@ -787,7 +787,7 @@
apply (auto simp add: divides_def simp del: pmult_Cons)
apply (rule_tac x = qb in exI)
apply (subgoal_tac "poly (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa)) =
- poly (pmult (pexp [uminus a, one] (?order a p)) (pmult [uminus a, one] qb))")
+ poly (pmult (pexp [uminus a, one] (order a p)) (pmult [uminus a, one] qb))")
apply (drule poly_mult_left_cancel [THEN iffD1], force)
apply (subgoal_tac "poly (pmult (pexp [uminus a, one] (order a q))
(pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa))) =
--- a/src/HOL/Decision_Procs/approximation.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Decision_Procs/approximation.ML Mon Mar 30 11:36:30 2015 +0200
@@ -5,11 +5,103 @@
signature APPROXIMATION =
sig
val approx: int -> Proof.context -> term -> term
+ val approximate: Proof.context -> term -> term
+ val approximation_tac : int -> (string * int) list -> int option -> Proof.context -> int -> tactic
end
structure Approximation: APPROXIMATION =
struct
+fun reorder_bounds_tac prems i =
+ let
+ fun variable_of_bound (Const (@{const_name Trueprop}, _) $
+ (Const (@{const_name Set.member}, _) $
+ Free (name, _) $ _)) = name
+ | variable_of_bound (Const (@{const_name Trueprop}, _) $
+ (Const (@{const_name HOL.eq}, _) $
+ Free (name, _) $ _)) = name
+ | variable_of_bound t = raise TERM ("variable_of_bound", [t])
+
+ val variable_bounds
+ = map (`(variable_of_bound o Thm.prop_of)) prems
+
+ fun add_deps (name, bnds)
+ = Graph.add_deps_acyclic (name,
+ remove (op =) name (Term.add_free_names (Thm.prop_of bnds) []))
+
+ val order = Graph.empty
+ |> fold Graph.new_node variable_bounds
+ |> fold add_deps variable_bounds
+ |> Graph.strong_conn |> map the_single |> rev
+ |> map_filter (AList.lookup (op =) variable_bounds)
+
+ fun prepend_prem th tac
+ = tac THEN rtac (th RSN (2, @{thm mp})) i
+ in
+ fold prepend_prem order all_tac
+ end
+
+fun approximation_conv ctxt ct =
+ approximation_oracle (Proof_Context.theory_of ctxt, Thm.term_of ct |> tap (tracing o Syntax.string_of_term ctxt));
+
+fun approximate ctxt t =
+ approximation_oracle (Proof_Context.theory_of ctxt, t)
+ |> Thm.prop_of |> Logic.dest_equals |> snd;
+
+(* Should be in HOL.thy ? *)
+fun gen_eval_tac conv ctxt = CONVERSION
+ (Object_Logic.judgment_conv (Conv.params_conv (~1) (K (Conv.concl_conv (~1) conv)) ctxt))
+ THEN' rtac TrueI
+
+val form_equations = @{thms interpret_form_equations};
+
+fun rewrite_interpret_form_tac ctxt prec splitting taylor i st = let
+ fun lookup_splitting (Free (name, _))
+ = case AList.lookup (op =) splitting name
+ of SOME s => HOLogic.mk_number @{typ nat} s
+ | NONE => @{term "0 :: nat"}
+ val vs = nth (Thm.prems_of st) (i - 1)
+ |> Logic.strip_imp_concl
+ |> HOLogic.dest_Trueprop
+ |> Term.strip_comb |> snd |> List.last
+ |> HOLogic.dest_list
+ val p = prec
+ |> HOLogic.mk_number @{typ nat}
+ |> Thm.cterm_of ctxt
+ in case taylor
+ of NONE => let
+ val n = vs |> length
+ |> HOLogic.mk_number @{typ nat}
+ |> Thm.cterm_of ctxt
+ val s = vs
+ |> map lookup_splitting
+ |> HOLogic.mk_list @{typ nat}
+ |> Thm.cterm_of ctxt
+ in
+ (rtac (Thm.instantiate ([], [(@{cpat "?n::nat"}, n),
+ (@{cpat "?prec::nat"}, p),
+ (@{cpat "?ss::nat list"}, s)])
+ @{thm "approx_form"}) i
+ THEN simp_tac (put_simpset (simpset_of @{context}) ctxt) i) st
+ end
+
+ | SOME t =>
+ if length vs <> 1
+ then raise (TERM ("More than one variable used for taylor series expansion", [Thm.prop_of st]))
+ else let
+ val t = t
+ |> HOLogic.mk_number @{typ nat}
+ |> Thm.cterm_of ctxt
+ val s = vs |> map lookup_splitting |> hd
+ |> Thm.cterm_of ctxt
+ in
+ rtac (Thm.instantiate ([], [(@{cpat "?s::nat"}, s),
+ (@{cpat "?t::nat"}, t),
+ (@{cpat "?prec::nat"}, p)])
+ @{thm "approx_tse_form"}) i st
+ end
+ end
+
fun calculated_subterms (@{const Trueprop} $ t) = calculated_subterms t
| calculated_subterms (@{const HOL.implies} $ _ $ t) = calculated_subterms t
| calculated_subterms (@{term "op <= :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) = [t1, t2]
@@ -156,4 +248,14 @@
(opt_modes -- Parse.term
>> (fn (modes, t) => Toplevel.keep (approximate_cmd modes t)));
-end;
+fun approximation_tac prec splitting taylor ctxt i =
+ REPEAT (FIRST' [etac @{thm intervalE},
+ etac @{thm meta_eqE},
+ rtac @{thm impI}] i)
+ THEN Subgoal.FOCUS (fn {prems, ...} => reorder_bounds_tac prems i) ctxt i
+ THEN DETERM (TRY (filter_prems_tac ctxt (K false) i))
+ THEN DETERM (Reification.tac ctxt form_equations NONE i)
+ THEN rewrite_interpret_form_tac ctxt prec splitting taylor i
+ THEN gen_eval_tac (approximation_conv ctxt) ctxt i
+
+end;
\ No newline at end of file
--- a/src/HOL/Decision_Procs/approximation_generator.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Decision_Procs/approximation_generator.ML Mon Mar 30 11:36:30 2015 +0200
@@ -141,7 +141,7 @@
#> conv_term ctxt (rewrite_with ctxt postproc_form_eqs))
frees
in
- if approximate ctxt (mk_approx_form e ts) |> is_True
+ if Approximation.approximate ctxt (mk_approx_form e ts) |> is_True
then SOME (true, ts')
else (if genuine_only then NONE else SOME (false, ts'))
end
@@ -161,6 +161,8 @@
"\<not> \<not> q \<longleftrightarrow> q"
by auto}
+val form_equations = @{thms interpret_form_equations};
+
fun reify_goal ctxt t =
HOLogic.mk_not t
|> conv_term ctxt (rewrite_with ctxt preproc_form_eqs)
--- a/src/HOL/Divides.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Divides.thy Mon Mar 30 11:36:30 2015 +0200
@@ -18,7 +18,7 @@
subsection {* Abstract division in commutative semirings. *}
-class semiring_div = comm_semiring_1_cancel + no_zero_divisors + div +
+class semiring_div = semidom + div +
assumes mod_div_equality: "a div b * b + a mod b = a"
and div_by_0 [simp]: "a div 0 = 0"
and div_0 [simp]: "0 div a = 0"
@@ -445,10 +445,10 @@
end
-class ring_div = semiring_div + comm_ring_1
+class ring_div = comm_ring_1 + semiring_div
begin
-subclass ring_1_no_zero_divisors ..
+subclass idom ..
text {* Negation respects modular equivalence. *}
@@ -548,7 +548,7 @@
subsubsection {* Parity and division *}
-class semiring_div_parity = semiring_div + semiring_numeral +
+class semiring_div_parity = semiring_div + comm_semiring_1_diff_distrib + numeral +
assumes parity: "a mod 2 = 0 \<or> a mod 2 = 1"
assumes one_mod_two_eq_one [simp]: "1 mod 2 = 1"
assumes zero_not_eq_two: "0 \<noteq> 2"
@@ -583,19 +583,6 @@
subclass semiring_parity
proof (unfold_locales, unfold dvd_eq_mod_eq_0 not_mod_2_eq_0_eq_1)
- fix a b c
- show "(c * a + b) mod a = 0 \<longleftrightarrow> b mod a = 0"
- by simp
-next
- fix a b c
- assume "(b + c) mod a = 0"
- with mod_add_eq [of b c a]
- have "(b mod a + c mod a) mod a = 0"
- by simp
- moreover assume "b mod a = 0"
- ultimately show "c mod a = 0"
- by simp
-next
show "1 mod 2 = 1"
by (fact one_mod_two_eq_one)
next
@@ -651,11 +638,9 @@
and less technical class hierarchy.
*}
-class semiring_numeral_div = linordered_semidom + minus + semiring_div +
- assumes diff_invert_add1: "a + b = c \<Longrightarrow> a = c - b"
- and le_add_diff_inverse2: "b \<le> a \<Longrightarrow> a - b + b = a"
- assumes mult_div_cancel: "b * (a div b) = a - a mod b"
- and div_less: "0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a div b = 0"
+class semiring_numeral_div = semiring_div + comm_semiring_1_diff_distrib + linordered_semidom +
+ assumes le_add_diff_inverse2: "b \<le> a \<Longrightarrow> a - b + b = a"
+ assumes div_less: "0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a div b = 0"
and mod_less: " 0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a mod b = a"
and div_positive: "0 < b \<Longrightarrow> b \<le> a \<Longrightarrow> a div b > 0"
and mod_less_eq_dividend: "0 \<le> a \<Longrightarrow> a mod b \<le> a"
@@ -666,9 +651,16 @@
assumes discrete: "a < b \<longleftrightarrow> a + 1 \<le> b"
begin
-lemma diff_zero [simp]:
- "a - 0 = a"
- by (rule diff_invert_add1 [symmetric]) simp
+lemma mult_div_cancel:
+ "b * (a div b) = a - a mod b"
+proof -
+ have "b * (a div b) + a mod b = a"
+ using mod_div_equality [of a b] by (simp add: ac_simps)
+ then have "b * (a div b) + a mod b - a mod b = a - a mod b"
+ by simp
+ then show ?thesis
+ by simp
+qed
subclass semiring_div_parity
proof
@@ -713,7 +705,7 @@
by (simp add: w_def div_mult2_eq mult_div_cancel ac_simps)
with `w = 1` have div: "2 * (a div (2 * b)) = a div b - 1" by simp
then show ?P and ?Q
- by (simp_all add: div mod diff_invert_add1 [symmetric] le_add_diff_inverse2)
+ by (simp_all add: div mod add_implies_diff [symmetric] le_add_diff_inverse2)
qed
lemma divmod_digit_0:
@@ -862,7 +854,7 @@
end
-hide_fact (open) diff_invert_add1 le_add_diff_inverse2 diff_zero
+hide_fact (open) le_add_diff_inverse2
-- {* restore simple accesses for more general variants of theorems *}
@@ -1195,7 +1187,7 @@
lemma mod_lemma: "[| (0::nat) < c; r < b |] ==> b * (q mod c) + r < b * c"
apply (cut_tac m = q and n = c in mod_less_divisor)
apply (drule_tac [2] m = "q mod c" in less_imp_Suc_add, auto)
- apply (erule_tac P = "%x. ?lhs < ?rhs x" in ssubst)
+ apply (erule_tac P = "%x. lhs < rhs x" for lhs rhs in ssubst)
apply (simp add: add_mult_distrib2)
done
@@ -2408,7 +2400,7 @@
"0<k ==>
P(n div k :: int)(n mod k) = (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P i j)"
apply (rule iffI, clarify)
- apply (erule_tac P="P ?x ?y" in rev_mp)
+ apply (erule_tac P="P x y" for x y in rev_mp)
apply (subst mod_add_eq)
apply (subst zdiv_zadd1_eq)
apply (simp add: div_pos_pos_trivial mod_pos_pos_trivial)
@@ -2421,7 +2413,7 @@
"k<0 ==>
P(n div k :: int)(n mod k) = (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P i j)"
apply (rule iffI, clarify)
- apply (erule_tac P="P ?x ?y" in rev_mp)
+ apply (erule_tac P="P x y" for x y in rev_mp)
apply (subst mod_add_eq)
apply (subst zdiv_zadd1_eq)
apply (simp add: div_neg_neg_trivial mod_neg_neg_trivial)
@@ -2857,4 +2849,3 @@
code_module Divides \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
end
-
--- a/src/HOL/GCD.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/GCD.thy Mon Mar 30 11:36:30 2015 +0200
@@ -626,11 +626,11 @@
apply (erule subst)
apply (rule iffI)
apply force
- apply (drule_tac x = "abs ?e" in exI)
- apply (case_tac "(?e::int) >= 0")
+ apply (drule_tac x = "abs e" for e in exI)
+ apply (case_tac "e >= 0" for e :: int)
apply force
apply force
-done
+ done
lemma gcd_coprime_nat:
assumes z: "gcd (a::nat) b \<noteq> 0" and a: "a = a' * gcd a b" and
--- a/src/HOL/Groups.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Groups.thy Mon Mar 30 11:36:30 2015 +0200
@@ -248,55 +248,52 @@
end
-class cancel_ab_semigroup_add = ab_semigroup_add +
- assumes add_imp_eq: "a + b = a + c \<Longrightarrow> b = c"
+class cancel_ab_semigroup_add = ab_semigroup_add + minus +
+ assumes add_diff_cancel_left' [simp]: "(a + b) - a = b"
+ assumes diff_diff_add [algebra_simps, field_simps]: "a - b - c = a - (b + c)"
begin
+lemma add_diff_cancel_right' [simp]:
+ "(a + b) - b = a"
+ using add_diff_cancel_left' [of b a] by (simp add: ac_simps)
+
subclass cancel_semigroup_add
proof
fix a b c :: 'a
- assume "a + b = a + c"
- then show "b = c" by (rule add_imp_eq)
+ assume "a + b = a + c"
+ then have "a + b - a = a + c - a"
+ by simp
+ then show "b = c"
+ by simp
next
fix a b c :: 'a
assume "b + a = c + a"
- then have "a + b = a + c" by (simp only: add.commute)
- then show "b = c" by (rule add_imp_eq)
+ then have "b + a - a = c + a - a"
+ by simp
+ then show "b = c"
+ by simp
qed
+lemma add_diff_cancel_left [simp]:
+ "(c + a) - (c + b) = a - b"
+ unfolding diff_diff_add [symmetric] by simp
+
+lemma add_diff_cancel_right [simp]:
+ "(a + c) - (b + c) = a - b"
+ using add_diff_cancel_left [symmetric] by (simp add: ac_simps)
+
+lemma diff_right_commute:
+ "a - c - b = a - b - c"
+ by (simp add: diff_diff_add add.commute)
+
end
class cancel_comm_monoid_add = cancel_ab_semigroup_add + comm_monoid_add
-
-class comm_monoid_diff = comm_monoid_add + minus +
- assumes diff_zero [simp]: "a - 0 = a"
- and zero_diff [simp]: "0 - a = 0"
- and add_diff_cancel_left [simp]: "(c + a) - (c + b) = a - b"
- and diff_diff_add: "a - b - c = a - (b + c)"
begin
-lemma add_diff_cancel_right [simp]:
- "(a + c) - (b + c) = a - b"
- using add_diff_cancel_left [symmetric] by (simp add: add.commute)
-
-lemma add_diff_cancel_left' [simp]:
- "(b + a) - b = a"
-proof -
- have "(b + a) - (b + 0) = a" by (simp only: add_diff_cancel_left diff_zero)
- then show ?thesis by simp
-qed
-
-lemma add_diff_cancel_right' [simp]:
- "(a + b) - b = a"
- using add_diff_cancel_left' [symmetric] by (simp add: add.commute)
-
-lemma diff_add_zero [simp]:
- "a - (a + b) = 0"
-proof -
- have "a - (a + b) = (a + 0) - (a + b)" by simp
- also have "\<dots> = 0" by (simp only: add_diff_cancel_left zero_diff)
- finally show ?thesis .
-qed
+lemma diff_zero [simp]:
+ "a - 0 = a"
+ using add_diff_cancel_right' [of a 0] by simp
lemma diff_cancel [simp]:
"a - a = 0"
@@ -305,10 +302,6 @@
then show ?thesis by simp
qed
-lemma diff_right_commute:
- "a - c - b = a - b - c"
- by (simp add: diff_diff_add add.commute)
-
lemma add_implies_diff:
assumes "c + b = a"
shows "c = a - b"
@@ -317,6 +310,20 @@
then show "c = a - b" by simp
qed
+end
+
+class comm_monoid_diff = cancel_comm_monoid_add +
+ assumes zero_diff [simp]: "0 - a = 0"
+begin
+
+lemma diff_add_zero [simp]:
+ "a - (a + b) = 0"
+proof -
+ have "a - (a + b) = (a + 0) - (a + b)" by simp
+ also have "\<dots> = 0" by (simp only: add_diff_cancel_left zero_diff)
+ finally show ?thesis .
+qed
+
end
@@ -527,10 +534,12 @@
subclass cancel_comm_monoid_add
proof
fix a b c :: 'a
- assume "a + b = a + c"
- then have "- a + a + b = - a + a + c"
- by (simp only: add.assoc)
- then show "b = c" by simp
+ have "b + a - a = b"
+ by simp
+ then show "a + b - a = b"
+ by (simp add: ac_simps)
+ show "a - b - c = a - (b + c)"
+ by (simp add: algebra_simps)
qed
lemma uminus_add_conv_diff [simp]:
@@ -545,18 +554,6 @@
"(a - b) + c = (a + c) - b"
by (simp add: algebra_simps)
-lemma diff_diff_eq [algebra_simps, field_simps]:
- "(a - b) - c = a - (b + c)"
- by (simp add: algebra_simps)
-
-lemma diff_add_eq_diff_diff:
- "a - (b + c) = a - b - c"
- using diff_add_eq_diff_diff_swap [of a c b] by (simp add: add.commute)
-
-lemma add_diff_cancel_left [simp]:
- "(c + a) - (c + b) = a - b"
- by (simp add: algebra_simps)
-
end
@@ -1375,7 +1372,13 @@
end
-hide_fact (open) ab_diff_conv_add_uminus add_imp_eq
+hide_fact (open) ab_diff_conv_add_uminus add_0 mult_1 ab_left_minus
+
+lemmas add_0 = add_0_left -- \<open>FIXME duplicate\<close>
+lemmas mult_1 = mult_1_left -- \<open>FIXME duplicate\<close>
+lemmas ab_left_minus = left_minus -- \<open>FIXME duplicate\<close>
+lemmas diff_diff_eq = diff_diff_add -- \<open>FIXME duplicate\<close>
+
subsection {* Tools setup *}
--- a/src/HOL/Groups_Big.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Groups_Big.thy Mon Mar 30 11:36:30 2015 +0200
@@ -1150,7 +1150,7 @@
lemma setprod_zero_iff [simp]:
assumes "finite A"
- shows "setprod f A = (0::'a::{comm_semiring_1,no_zero_divisors}) \<longleftrightarrow> (\<exists>a\<in>A. f a = 0)"
+ shows "setprod f A = (0::'a::semidom) \<longleftrightarrow> (\<exists>a\<in>A. f a = 0)"
using assms by (induct A) (auto simp: no_zero_divisors)
lemma (in field) setprod_diff1:
--- a/src/HOL/HOLCF/IOA/meta_theory/Automata.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/HOLCF/IOA/meta_theory/Automata.thy Mon Mar 30 11:36:30 2015 +0200
@@ -654,7 +654,7 @@
apply (simp add: is_asig_of_def rename_def rename_set_def asig_internals_def
asig_outputs_def asig_inputs_def actions_def is_asig_def asig_of_def)
apply auto
-apply (drule_tac [!] s = "Some ?x" in sym)
+apply (drule_tac [!] s = "Some _" in sym)
apply auto
done
--- a/src/HOL/HOLCF/Tools/Domain/domain_induction.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/HOLCF/Tools/Domain/domain_induction.ML Mon Mar 30 11:36:30 2015 +0200
@@ -176,7 +176,7 @@
val tacs1 = [
quant_tac ctxt 1,
simp_tac (put_simpset HOL_ss ctxt) 1,
- Induct_Tacs.induct_tac ctxt [[SOME "n"]] 1,
+ Induct_Tacs.induct_tac ctxt [[SOME "n"]] NONE 1,
simp_tac (put_simpset take_ss ctxt addsimps prems) 1,
TRY (safe_tac (put_claset HOL_cs ctxt))]
fun con_tac _ =
--- a/src/HOL/Hoare_Parallel/Graph.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Hoare_Parallel/Graph.thy Mon Mar 30 11:36:30 2015 +0200
@@ -287,10 +287,10 @@
apply(case_tac "length path")
apply force
apply simp
- apply(erule_tac P = "\<lambda>i. i < nata \<longrightarrow> ?P i" and x = "nat" in allE)
+ apply(erule_tac P = "\<lambda>i. i < nata \<longrightarrow> P i" and x = "nat" for P in allE)
apply simp
apply clarify
- apply(erule_tac P = "\<lambda>i. i < Suc nat \<longrightarrow> ?P i" and x = "nat" in allE)
+ apply(erule_tac P = "\<lambda>i. i < Suc nat \<longrightarrow> P i" and x = "nat" for P in allE)
apply simp
apply(case_tac "j<I")
apply(erule_tac x = "j" in allE)
@@ -329,10 +329,10 @@
apply(case_tac "length path")
apply force
apply simp
- apply(erule_tac P = "\<lambda>i. i < nata \<longrightarrow> ?P i" and x = "nat" in allE)
+ apply(erule_tac P = "\<lambda>i. i < nata \<longrightarrow> P i" and x = "nat" for P in allE)
apply simp
apply clarify
- apply(erule_tac P = "\<lambda>i. i < Suc nat \<longrightarrow> ?P i" and x = "nat" in allE)
+ apply(erule_tac P = "\<lambda>i. i < Suc nat \<longrightarrow> P i" and x = "nat" for P in allE)
apply simp
apply(case_tac "j\<le>R")
apply(drule le_imp_less_or_eq [of _ R])
--- a/src/HOL/Hoare_Parallel/OG_Hoare.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Hoare_Parallel/OG_Hoare.thy Mon Mar 30 11:36:30 2015 +0200
@@ -441,7 +441,7 @@
apply simp
apply force
apply simp
- apply(erule_tac V = "\<forall>i. ?P i" in thin_rl)
+ apply(erule_tac V = "\<forall>i. P i" for P in thin_rl)
apply(drule_tac s = "length Rs" in sym)
apply(erule allE, erule impE, assumption)
apply(force dest: nth_mem simp add: All_None_def)
--- a/src/HOL/Hoare_Parallel/RG_Hoare.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Hoare_Parallel/RG_Hoare.thy Mon Mar 30 11:36:30 2015 +0200
@@ -186,7 +186,7 @@
apply force
apply clarify
apply(erule_tac x="Suc ia" in allE,simp)
-apply(erule_tac x="0" and P="\<lambda>j. ?H j \<longrightarrow> (?J j) \<notin> ctran" in allE,simp)
+apply(erule_tac x="0" and P="\<lambda>j. H j \<longrightarrow> (J j) \<notin> ctran" for H J in allE,simp)
done
lemma etran_or_ctran2 [rule_format]:
@@ -241,14 +241,14 @@
apply(case_tac k,simp,simp)
apply(case_tac j,simp)
apply(erule_tac x=0 in allE)
- apply(erule_tac x="nat" and P="\<lambda>j. (0\<le>j) \<longrightarrow> (?J j)" in allE,simp)
+ apply(erule_tac x="nat" and P="\<lambda>j. (0\<le>j) \<longrightarrow> (J j)" for J in allE,simp)
apply(subgoal_tac "t\<in>p")
apply(subgoal_tac "(\<forall>i. i < length xs \<longrightarrow> ((P, t) # xs) ! i -e\<rightarrow> xs ! i \<longrightarrow> (snd (((P, t) # xs) ! i), snd (xs ! i)) \<in> rely)")
apply clarify
- apply(erule_tac x="Suc i" and P="\<lambda>j. (?H j) \<longrightarrow> (?J j)\<in>etran" in allE,simp)
+ apply(erule_tac x="Suc i" and P="\<lambda>j. (H j) \<longrightarrow> (J j)\<in>etran" for H J in allE,simp)
apply clarify
- apply(erule_tac x="Suc i" and P="\<lambda>j. (?H j) \<longrightarrow> (?J j) \<longrightarrow> (?T j)\<in>rely" in allE,simp)
- apply(erule_tac x=0 and P="\<lambda>j. (?H j) \<longrightarrow> (?J j)\<in>etran \<longrightarrow> ?T j" in allE,simp)
+ apply(erule_tac x="Suc i" and P="\<lambda>j. (H j) \<longrightarrow> (J j) \<longrightarrow> (T j)\<in>rely" for H J T in allE,simp)
+ apply(erule_tac x=0 and P="\<lambda>j. (H j) \<longrightarrow> (J j)\<in>etran \<longrightarrow> T j" for H J T in allE,simp)
apply(simp(no_asm_use) only:stable_def)
apply(erule_tac x=s in allE)
apply(erule_tac x=t in allE)
@@ -258,23 +258,23 @@
apply(rule Env)
apply simp
apply(erule_tac x="nata" in allE)
- apply(erule_tac x="nat" and P="\<lambda>j. (?s\<le>j) \<longrightarrow> (?J j)" in allE,simp)
+ apply(erule_tac x="nat" and P="\<lambda>j. (s\<le>j) \<longrightarrow> (J j)" for s J in allE,simp)
apply(subgoal_tac "(\<forall>i. i < length xs \<longrightarrow> ((P, t) # xs) ! i -e\<rightarrow> xs ! i \<longrightarrow> (snd (((P, t) # xs) ! i), snd (xs ! i)) \<in> rely)")
apply clarify
- apply(erule_tac x="Suc i" and P="\<lambda>j. (?H j) \<longrightarrow> (?J j)\<in>etran" in allE,simp)
+ apply(erule_tac x="Suc i" and P="\<lambda>j. (H j) \<longrightarrow> (J j)\<in>etran" for H J in allE,simp)
apply clarify
- apply(erule_tac x="Suc i" and P="\<lambda>j. (?H j) \<longrightarrow> (?J j) \<longrightarrow> (?T j)\<in>rely" in allE,simp)
+ apply(erule_tac x="Suc i" and P="\<lambda>j. (H j) \<longrightarrow> (J j) \<longrightarrow> (T j)\<in>rely" for H J T in allE,simp)
apply(case_tac k,simp,simp)
apply(case_tac j)
- apply(erule_tac x=0 and P="\<lambda>j. (?H j) \<longrightarrow> (?J j)\<in>etran" in allE,simp)
+ apply(erule_tac x=0 and P="\<lambda>j. (H j) \<longrightarrow> (J j)\<in>etran" for H J in allE,simp)
apply(erule etran.cases,simp)
apply(erule_tac x="nata" in allE)
-apply(erule_tac x="nat" and P="\<lambda>j. (?s\<le>j) \<longrightarrow> (?J j)" in allE,simp)
+apply(erule_tac x="nat" and P="\<lambda>j. (s\<le>j) \<longrightarrow> (J j)" for s J in allE,simp)
apply(subgoal_tac "(\<forall>i. i < length xs \<longrightarrow> ((Q, t) # xs) ! i -e\<rightarrow> xs ! i \<longrightarrow> (snd (((Q, t) # xs) ! i), snd (xs ! i)) \<in> rely)")
apply clarify
- apply(erule_tac x="Suc i" and P="\<lambda>j. (?H j) \<longrightarrow> (?J j)\<in>etran" in allE,simp)
+ apply(erule_tac x="Suc i" and P="\<lambda>j. (H j) \<longrightarrow> (J j)\<in>etran" for H J in allE,simp)
apply clarify
-apply(erule_tac x="Suc i" and P="\<lambda>j. (?H j) \<longrightarrow> (?J j) \<longrightarrow> (?T j)\<in>rely" in allE,simp)
+apply(erule_tac x="Suc i" and P="\<lambda>j. (H j) \<longrightarrow> (J j) \<longrightarrow> (T j)\<in>rely" for H J T in allE,simp)
done
subsection \<open>Soundness of the System for Component Programs\<close>
@@ -341,7 +341,7 @@
apply(case_tac "x!i")
apply clarify
apply(drule_tac s="Some (Basic f)" in sym,simp)
- apply(thin_tac "\<forall>j. ?H j")
+ apply(thin_tac "\<forall>j. H j" for H)
apply(force elim:ctran.cases)
apply clarify
apply(simp add:cp_def)
@@ -468,7 +468,7 @@
apply(drule_tac c=l in subsetD)
apply (simp add:cp_def)
apply clarify
- apply(erule_tac x=ia and P="\<lambda>i. ?H i \<longrightarrow> (?J i,?I i)\<in>ctran" in allE,simp)
+ apply(erule_tac x=ia and P="\<lambda>i. H i \<longrightarrow> (J i, I i)\<in>ctran" for H J I in allE,simp)
apply(erule etranE,simp)
apply simp
apply clarify
@@ -499,7 +499,7 @@
apply(drule_tac c=l in subsetD)
apply (simp add:cp_def)
apply clarify
- apply(erule_tac x=i and P="\<lambda>i. ?H i \<longrightarrow> (?J i,?I i)\<in>ctran" in allE,simp)
+ apply(erule_tac x=i and P="\<lambda>i. H i \<longrightarrow> (J i, I i)\<in>ctran" for H J I in allE,simp)
apply(erule etranE,simp)
apply simp
apply clarify
@@ -532,7 +532,7 @@
apply(erule_tac m="length x" in etran_or_ctran,simp+)
apply(case_tac x, (simp add:last_length)+)
apply(erule exE)
-apply(drule_tac n=i and P="\<lambda>i. ?H i \<and> (?J i,?I i)\<in> ctran" in Ex_first_occurrence)
+apply(drule_tac n=i and P="\<lambda>i. H i \<and> (J i, I i) \<in> ctran" for H J I in Ex_first_occurrence)
apply clarify
apply (simp add:assum_def)
apply(frule_tac j=0 and k="m" and p=pre in stability,simp+)
@@ -549,7 +549,7 @@
apply(erule disjE)
apply(erule_tac x=i in allE, erule impE, assumption)
apply simp+
- apply(erule_tac x="i - (Suc m)" and P="\<lambda>j. ?H j \<longrightarrow> ?J j \<longrightarrow> (?I j)\<in>guar" in allE)
+ apply(erule_tac x="i - (Suc m)" and P="\<lambda>j. H j \<longrightarrow> J j \<longrightarrow> (I j)\<in>guar" for H J I in allE)
apply(subgoal_tac "(Suc m)+(i - Suc m) \<le> length x")
apply(subgoal_tac "(Suc m)+Suc (i - Suc m) \<le> length x")
apply(rotate_tac -2)
@@ -569,7 +569,7 @@
apply(erule_tac x=i in allE, erule impE, assumption)
apply simp
apply simp
-apply(erule_tac x="i - (Suc m)" and P="\<lambda>j. ?H j \<longrightarrow> ?J j \<longrightarrow> (?I j)\<in>guar" in allE)
+apply(erule_tac x="i - (Suc m)" and P="\<lambda>j. H j \<longrightarrow> J j \<longrightarrow> (I j)\<in>guar" for H J I in allE)
apply(subgoal_tac "(Suc m)+(i - Suc m) \<le> length x")
apply(subgoal_tac "(Suc m)+Suc (i - Suc m) \<le> length x")
apply(rotate_tac -2)
@@ -713,14 +713,14 @@
apply(drule_tac c=xs in subsetD,simp add:cp_def cptn_iff_cptn_mod)
apply(simp add:assum_def)
apply clarify
- apply(erule_tac P="\<lambda>j. ?H j \<longrightarrow> ?J j \<longrightarrow> ?I j" in allE,erule impE, assumption)
+ apply(erule_tac P="\<lambda>j. H j \<longrightarrow> J j \<longrightarrow> I j" for H J I in allE,erule impE, assumption)
apply(simp add:snd_lift)
apply(erule mp)
apply(force elim:etranE intro:Env simp add:lift_def)
apply(simp add:comm_def)
apply(rule conjI)
apply clarify
- apply(erule_tac P="\<lambda>j. ?H j \<longrightarrow> ?J j \<longrightarrow> ?I j" in allE,erule impE, assumption)
+ apply(erule_tac P="\<lambda>j. H j \<longrightarrow> J j \<longrightarrow> I j" for H J I in allE,erule impE, assumption)
apply(simp add:snd_lift)
apply(erule mp)
apply(case_tac "(xs!i)")
@@ -914,7 +914,7 @@
apply clarify
apply (simp add:last_length)
--\<open>WhileOne\<close>
-apply(thin_tac "P = While b P \<longrightarrow> ?Q")
+apply(thin_tac "P = While b P \<longrightarrow> Q" for Q)
apply(rule ctran_in_comm,simp)
apply(simp add:Cons_lift del:list.map)
apply(simp add:comm_def del:list.map)
@@ -957,7 +957,7 @@
apply(simp only:last_lift_not_None)
apply simp
--\<open>WhileMore\<close>
-apply(thin_tac "P = While b P \<longrightarrow> ?Q")
+apply(thin_tac "P = While b P \<longrightarrow> Q" for Q)
apply(rule ctran_in_comm,simp del:last.simps)
--\<open>metiendo la hipotesis antes de dividir la conclusion.\<close>
apply(subgoal_tac "(Some (While b P), snd (last ((Some P, sa) # xs))) # ys \<in> assum (pre, rely)")
@@ -979,7 +979,7 @@
apply simp
apply simp
apply(simp add:snd_lift del:list.map last.simps)
- apply(thin_tac " \<forall>i. i < length ys \<longrightarrow> ?P i")
+ apply(thin_tac " \<forall>i. i < length ys \<longrightarrow> P i" for P)
apply(simp only:com_validity_def cp_def cptn_iff_cptn_mod)
apply(erule_tac x=sa in allE)
apply(drule_tac c="(Some P, sa) # xs" in subsetD)
@@ -1111,13 +1111,13 @@
apply simp
apply(erule exE)
--\<open>the first c-tran that does not satisfy the guarantee-condition is from @{text "\<sigma>_i"} at step @{text "m"}.\<close>
-apply(drule_tac n=j and P="\<lambda>j. \<exists>i. ?H i j" in Ex_first_occurrence)
+apply(drule_tac n=j and P="\<lambda>j. \<exists>i. H i j" for H in Ex_first_occurrence)
apply(erule exE)
apply clarify
--\<open>@{text "\<sigma>_i \<in> A(pre, rely_1)"}\<close>
apply(subgoal_tac "take (Suc (Suc m)) (clist!i) \<in> assum(Pre(xs!i), Rely(xs!i))")
--\<open>but this contradicts @{text "\<Turnstile> \<sigma>_i sat [pre_i,rely_i,guar_i,post_i]"}\<close>
- apply(erule_tac x=i and P="\<lambda>i. ?H i \<longrightarrow> \<Turnstile> (?J i) sat [?I i,?K i,?M i,?N i]" in allE,erule impE,assumption)
+ apply(erule_tac x=i and P="\<lambda>i. H i \<longrightarrow> \<Turnstile> (J i) sat [I i,K i,M i,N i]" for H J I K M N in allE,erule impE,assumption)
apply(simp add:com_validity_def)
apply(erule_tac x=s in allE)
apply(simp add:cp_def comm_def)
@@ -1126,40 +1126,40 @@
apply (blast intro: takecptn_is_cptn)
apply simp
apply clarify
- apply(erule_tac x=m and P="\<lambda>j. ?I j \<and> ?J j \<longrightarrow> ?H j" in allE)
+ apply(erule_tac x=m and P="\<lambda>j. I j \<and> J j \<longrightarrow> H j" for I J H in allE)
apply (simp add:conjoin_def same_length_def)
apply(simp add:assum_def)
apply(rule conjI)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<in>cp (?K j) (?J j)" in allE)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> I j \<in>cp (K j) (J j)" for H I K J in allE)
apply(simp add:cp_def par_assum_def)
apply(drule_tac c="s" in subsetD,simp)
apply simp
apply clarify
-apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> ?M \<union> UNION (?S j) (?T j) \<subseteq> (?L j)" in allE)
+apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> M \<union> UNION (S j) (T j) \<subseteq> (L j)" for H M S T L in allE)
apply simp
apply(erule subsetD)
apply simp
apply(simp add:conjoin_def compat_label_def)
apply clarify
-apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> (?P j) \<or> ?Q j" in allE,simp)
+apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> (P j) \<or> Q j" for H P Q in allE,simp)
--\<open>each etran in @{text "\<sigma>_1[0\<dots>m]"} corresponds to\<close>
apply(erule disjE)
--\<open>a c-tran in some @{text "\<sigma>_{ib}"}\<close>
apply clarify
apply(case_tac "i=ib",simp)
apply(erule etranE,simp)
- apply(erule_tac x="ib" and P="\<lambda>i. ?H i \<longrightarrow> (?I i) \<or> (?J i)" in allE)
+ apply(erule_tac x="ib" and P="\<lambda>i. H i \<longrightarrow> (I i) \<or> (J i)" for H I J in allE)
apply (erule etranE)
apply(case_tac "ia=m",simp)
apply simp
- apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> (\<forall> i. ?P i j)" in allE)
+ apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> (\<forall>i. P i j)" for H P in allE)
apply(subgoal_tac "ia<m",simp)
prefer 2
apply arith
- apply(erule_tac x=ib and P="\<lambda>j. (?I j, ?H j)\<in> ctran \<longrightarrow> (?P i j)" in allE,simp)
+ apply(erule_tac x=ib and P="\<lambda>j. (I j, H j) \<in> ctran \<longrightarrow> P i j" for I H P in allE,simp)
apply(simp add:same_state_def)
- apply(erule_tac x=i and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in all_dupE)
- apply(erule_tac x=ib and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in allE,simp)
+ apply(erule_tac x=i and P="\<lambda>j. (T j) \<longrightarrow> (\<forall>i. (H j i) \<longrightarrow> (snd (d j i))=(snd (e j i)))" for T H d e in all_dupE)
+ apply(erule_tac x=ib and P="\<lambda>j. (T j) \<longrightarrow> (\<forall>i. (H j i) \<longrightarrow> (snd (d j i))=(snd (e j i)))" for T H d e in allE,simp)
--\<open>or an e-tran in @{text "\<sigma>"},
therefore it satisfies @{text "rely \<or> guar_{ib}"}\<close>
apply (force simp add:par_assum_def same_state_def)
@@ -1181,19 +1181,19 @@
apply clarify
apply(simp add:conjoin_def compat_label_def)
apply clarify
-apply(erule_tac x=j and P="\<lambda>j. ?H j \<longrightarrow> (?J j \<and> (\<exists>i. ?P i j)) \<or> ?I j" in allE,simp)
+apply(erule_tac x=j and P="\<lambda>j. H j \<longrightarrow> (J j \<and> (\<exists>i. P i j)) \<or> I j" for H J P I in allE,simp)
apply(erule disjE)
prefer 2
apply(force simp add:same_state_def par_assum_def)
apply clarify
apply(case_tac "i=ia",simp)
apply(erule etranE,simp)
-apply(erule_tac x="ia" and P="\<lambda>i. ?H i \<longrightarrow> (?I i) \<or> (?J i)" in allE,simp)
-apply(erule_tac x=j and P="\<lambda>j. \<forall>i. ?S j i \<longrightarrow> (?I j i, ?H j i)\<in> ctran \<longrightarrow> (?P i j)" in allE)
-apply(erule_tac x=ia and P="\<lambda>j. ?S j \<longrightarrow> (?I j, ?H j)\<in> ctran \<longrightarrow> (?P j)" in allE)
+apply(erule_tac x="ia" and P="\<lambda>i. H i \<longrightarrow> (I i) \<or> (J i)" for H I J in allE,simp)
+apply(erule_tac x=j and P="\<lambda>j. \<forall>i. S j i \<longrightarrow> (I j i, H j i) \<in> ctran \<longrightarrow> P i j" for S I H P in allE)
+apply(erule_tac x=ia and P="\<lambda>j. S j \<longrightarrow> (I j, H j) \<in> ctran \<longrightarrow> P j" for S I H P in allE)
apply(simp add:same_state_def)
-apply(erule_tac x=i and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in all_dupE)
-apply(erule_tac x=ia and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in allE,simp)
+apply(erule_tac x=i and P="\<lambda>j. T j \<longrightarrow> (\<forall>i. H j i \<longrightarrow> (snd (d j i))=(snd (e j i)))" for T H d e in all_dupE)
+apply(erule_tac x=ia and P="\<lambda>j. T j \<longrightarrow> (\<forall>i. H j i \<longrightarrow> (snd (d j i))=(snd (e j i)))" for T H d e in allE,simp)
done
lemma four:
@@ -1224,18 +1224,18 @@
apply assumption
apply(simp add:conjoin_def same_program_def)
apply clarify
-apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> fst(?I j)=(?J j)" in all_dupE)
-apply(erule_tac x="Suc i" and P="\<lambda>j. ?H j \<longrightarrow> fst(?I j)=(?J j)" in allE)
+apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> fst(I j)=(J j)" for H I J in all_dupE)
+apply(erule_tac x="Suc i" and P="\<lambda>j. H j \<longrightarrow> fst(I j)=(J j)" for H I J in allE)
apply(erule par_ctranE,simp)
-apply(erule_tac x=i and P="\<lambda>j. \<forall>i. ?S j i \<longrightarrow> (?I j i, ?H j i)\<in> ctran \<longrightarrow> (?P i j)" in allE)
-apply(erule_tac x=ia and P="\<lambda>j. ?S j \<longrightarrow> (?I j, ?H j)\<in> ctran \<longrightarrow> (?P j)" in allE)
+apply(erule_tac x=i and P="\<lambda>j. \<forall>i. S j i \<longrightarrow> (I j i, H j i) \<in> ctran \<longrightarrow> P i j" for S I H P in allE)
+apply(erule_tac x=ia and P="\<lambda>j. S j \<longrightarrow> (I j, H j) \<in> ctran \<longrightarrow> P j" for S I H P in allE)
apply(rule_tac x=ia in exI)
apply(simp add:same_state_def)
-apply(erule_tac x=ia and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in all_dupE,simp)
-apply(erule_tac x=ia and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in allE,simp)
-apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in all_dupE)
-apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in all_dupE,simp)
-apply(erule_tac x="Suc i" and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
+apply(erule_tac x=ia and P="\<lambda>j. T j \<longrightarrow> (\<forall>i. H j i \<longrightarrow> (snd (d j i))=(snd (e j i)))" for T H d e in all_dupE,simp)
+apply(erule_tac x=ia and P="\<lambda>j. T j \<longrightarrow> (\<forall>i. H j i \<longrightarrow> (snd (d j i))=(snd (e j i)))" for T H d e in allE,simp)
+apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in all_dupE)
+apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in all_dupE,simp)
+apply(erule_tac x="Suc i" and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE,simp)
apply(erule mp)
apply(subgoal_tac "r=fst(clist ! ia ! Suc i)",simp)
apply(drule_tac i=ia in list_eq_if)
@@ -1266,17 +1266,17 @@
apply simp
apply clarify
apply(subgoal_tac "\<forall>i<length clist. clist!i\<in>assum(Pre(xs!i), Rely(xs!i))")
- apply(erule_tac x=i and P="\<lambda>i. ?H i \<longrightarrow> \<Turnstile> (?J i) sat [?I i,?K i,?M i,?N i]" in allE,erule impE,assumption)
+ apply(erule_tac x=i and P="\<lambda>i. H i \<longrightarrow> \<Turnstile> (J i) sat [I i,K i,M i,N i]" for H J I K M N in allE,erule impE,assumption)
apply(simp add:com_validity_def)
apply(erule_tac x=s in allE)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (?I j) \<in> cp (?J j) s" in allE,simp)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (I j) \<in> cp (J j) s" for H I J in allE,simp)
apply(drule_tac c="clist!i" in subsetD)
apply (force simp add:Com_def)
apply(simp add:comm_def conjoin_def same_program_def del:last.simps)
apply clarify
- apply(erule_tac x="length x - 1" and P="\<lambda>j. ?H j \<longrightarrow> fst(?I j)=(?J j)" in allE)
+ apply(erule_tac x="length x - 1" and P="\<lambda>j. H j \<longrightarrow> fst(I j)=(J j)" for H I J in allE)
apply (simp add:All_None_def same_length_def)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> length(?J j)=(?K j)" in allE)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> length(J j)=(K j)" for H J K in allE)
apply(subgoal_tac "length x - 1 < length x",simp)
apply(case_tac "x\<noteq>[]")
apply(simp add: last_conv_nth)
@@ -1297,8 +1297,8 @@
apply(rule conjI)
apply(simp add:conjoin_def same_state_def par_cp_def)
apply clarify
- apply(erule_tac x=ia and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in allE,simp)
- apply(erule_tac x=0 and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
+ apply(erule_tac x=ia and P="\<lambda>j. T j \<longrightarrow> (\<forall>i. H j i \<longrightarrow> (snd (d j i))=(snd (e j i)))" for T H d e in allE,simp)
+ apply(erule_tac x=0 and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE)
apply(case_tac x,simp+)
apply (simp add:par_assum_def)
apply clarify
--- a/src/HOL/Hoare_Parallel/RG_Tran.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Hoare_Parallel/RG_Tran.thy Mon Mar 30 11:36:30 2015 +0200
@@ -494,7 +494,7 @@
apply(clarify)
apply(simp add:conjoin_def compat_label_def)
apply clarify
-apply(erule_tac x="0" and P="\<lambda>j. ?H j \<longrightarrow> (?P j \<or> ?Q j)" in all_dupE,simp)
+apply(erule_tac x="0" and P="\<lambda>j. H j \<longrightarrow> (P j \<or> Q j)" for H P Q in all_dupE, simp)
apply(erule disjE)
--\<open>first step is a Component step\<close>
apply clarify
@@ -504,14 +504,14 @@
prefer 2
apply(simp add: same_state_def)
apply(erule_tac x=i in allE,erule impE,assumption,
- erule_tac x=1 and P="\<lambda>j. (?H j) \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
+ erule_tac x=1 and P="\<lambda>j. (H j) \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE, simp)
prefer 2
apply(simp add:same_program_def)
- apply(erule_tac x=1 and P="\<lambda>j. ?H j \<longrightarrow> (fst (?s j))=(?t j)" in allE,simp)
+ apply(erule_tac x=1 and P="\<lambda>j. H j \<longrightarrow> (fst (s j))=(t j)" for H s t in allE,simp)
apply(rule nth_equalityI,simp)
apply clarify
apply(case_tac "i=ia",simp,simp)
- apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
+ apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE)
apply(drule_tac t=i in not_sym,simp)
apply(erule etranE,simp)
apply(rule ParCptnComp)
@@ -532,8 +532,8 @@
apply(force simp add:same_length_def length_Suc_conv)
apply(simp add:same_state_def)
apply(erule_tac x=ia in allE, erule impE, assumption,
- erule_tac x=1 and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
- apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
+ erule_tac x=1 and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE)
+ apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE)
apply(drule_tac t=i in not_sym,simp)
apply(erule etranE,simp)
apply(erule allE,erule impE,assumption,erule tl_in_cptn)
@@ -543,7 +543,7 @@
apply clarify
apply(case_tac j,simp,simp)
apply(erule_tac x=ia in allE, erule impE, assumption,
- erule_tac x="Suc(Suc nat)" and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
+ erule_tac x="Suc(Suc nat)" and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE,simp)
apply(force simp add:same_length_def length_Suc_conv)
apply(rule conjI)
apply(simp add:same_program_def)
@@ -552,11 +552,11 @@
apply(rule nth_equalityI,simp)
apply clarify
apply(case_tac "i=ia",simp,simp)
- apply(erule_tac x="Suc(Suc nat)" and P="\<lambda>j. ?H j \<longrightarrow> (fst (?s j))=(?t j)" in allE,simp)
+ apply(erule_tac x="Suc(Suc nat)" and P="\<lambda>j. H j \<longrightarrow> (fst (s j))=(t j)" for H s t in allE,simp)
apply(rule nth_equalityI,simp,simp)
apply(force simp add:length_Suc_conv)
apply(rule allI,rule impI)
- apply(erule_tac x="Suc j" and P="\<lambda>j. ?H j \<longrightarrow> (?I j \<or> ?J j)" in allE,simp)
+ apply(erule_tac x="Suc j" and P="\<lambda>j. H j \<longrightarrow> (I j \<or> J j)" for H I J in allE,simp)
apply(erule disjE)
apply clarify
apply(rule_tac x=ia in exI,simp)
@@ -564,13 +564,13 @@
apply(rule conjI)
apply(force simp add: length_Suc_conv)
apply clarify
- apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE,erule impE,assumption)
- apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE,erule impE,assumption)
+ apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE,erule impE,assumption)
+ apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE,erule impE,assumption)
apply simp
apply(case_tac j,simp)
apply(rule tl_zero)
apply(erule_tac x=l in allE, erule impE, assumption,
- erule_tac x=1 and P="\<lambda>j. (?H j) \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
+ erule_tac x=1 and P="\<lambda>j. (H j) \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE,simp)
apply(force elim:etranE intro:Env)
apply force
apply force
@@ -585,8 +585,8 @@
apply(rule nth_tl_if)
apply force
apply(erule_tac x=ia in allE, erule impE, assumption,
- erule_tac x=1 and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
- apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
+ erule_tac x=1 and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE)
+ apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE)
apply(drule_tac t=i in not_sym,simp)
apply(erule etranE,simp)
apply(erule tl_zero)
@@ -595,41 +595,41 @@
apply clarify
apply(case_tac "i=l",simp)
apply(rule nth_tl_if)
- apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
apply simp
- apply(erule_tac P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE,erule impE,assumption,erule impE,assumption)
+ apply(erule_tac P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE,erule impE,assumption,erule impE,assumption)
apply(erule tl_zero,force)
- apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
apply(rule nth_tl_if)
- apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
apply(erule_tac x=l in allE, erule impE, assumption,
- erule_tac x=1 and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
- apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE,erule impE, assumption,simp)
+ erule_tac x=1 and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE)
+ apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE,erule impE, assumption,simp)
apply(erule etranE,simp)
apply(rule tl_zero)
apply force
apply force
- apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
apply(rule disjI2)
apply(case_tac j,simp)
apply clarify
apply(rule tl_zero)
- apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> ?I j\<in>etran" in allE,erule impE, assumption)
+ apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> I j\<in>etran" for H I in allE,erule impE, assumption)
apply(case_tac "i=ia",simp,simp)
apply(erule_tac x=ia in allE, erule impE, assumption,
- erule_tac x=1 and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
- apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE,erule impE, assumption,simp)
+ erule_tac x=1 and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE)
+ apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE,erule impE, assumption,simp)
apply(force elim:etranE intro:Env)
apply force
- apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
apply simp
apply clarify
apply(rule tl_zero)
apply(rule tl_zero,force)
apply force
- apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
apply force
- apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
--\<open>first step is an environmental step\<close>
apply clarify
apply(erule par_etran.cases)
@@ -641,24 +641,24 @@
apply(rule_tac x="map tl clist" in exI,simp)
apply(rule conjI)
apply clarify
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (?I ?s j) \<in> cptn" in allE,simp)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> I j \<in> cptn" for H I in allE,simp)
apply(erule cptn.cases)
apply(simp add:same_length_def)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
apply(simp add:same_state_def)
apply(erule_tac x=i in allE, erule impE, assumption,
- erule_tac x=1 and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> ?J j \<in>etran" in allE,simp)
+ erule_tac x=1 and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE,simp)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> J j \<in>etran" for H J in allE,simp)
apply(erule etranE,simp)
apply(simp add:same_state_def same_length_def)
apply(rule conjI,clarify)
apply(case_tac j,simp,simp)
apply(erule_tac x=i in allE, erule impE, assumption,
- erule_tac x="Suc(Suc nat)" and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
+ erule_tac x="Suc(Suc nat)" and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE,simp)
apply(rule tl_zero)
apply(simp)
apply force
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
apply(rule conjI)
apply(simp add:same_program_def)
apply clarify
@@ -666,52 +666,52 @@
apply(rule nth_equalityI,simp)
apply clarify
apply simp
- apply(erule_tac x="Suc(Suc nat)" and P="\<lambda>j. ?H j \<longrightarrow> (fst (?s j))=(?t j)" in allE,simp)
+ apply(erule_tac x="Suc(Suc nat)" and P="\<lambda>j. H j \<longrightarrow> (fst (s j))=(t j)" for H s t in allE,simp)
apply(rule nth_equalityI,simp,simp)
apply(force simp add:length_Suc_conv)
apply(rule allI,rule impI)
-apply(erule_tac x="Suc j" and P="\<lambda>j. ?H j \<longrightarrow> (?I j \<or> ?J j)" in allE,simp)
+apply(erule_tac x="Suc j" and P="\<lambda>j. H j \<longrightarrow> (I j \<or> J j)" for H I J in allE,simp)
apply(erule disjE)
apply clarify
apply(rule_tac x=i in exI,simp)
apply(rule conjI)
- apply(erule_tac x=i and P="\<lambda>i. ?H i \<longrightarrow> ?J i \<in>etran" in allE, erule impE, assumption)
+ apply(erule_tac x=i and P="\<lambda>i. H i \<longrightarrow> J i \<in>etran" for H J in allE, erule impE, assumption)
apply(erule etranE,simp)
apply(erule_tac x=i in allE, erule impE, assumption,
- erule_tac x=1 and P="\<lambda>j. (?H j) \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
+ erule_tac x=1 and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE,simp)
apply(rule nth_tl_if)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
apply simp
apply(erule tl_zero,force)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
apply clarify
- apply(erule_tac x=l and P="\<lambda>i. ?H i \<longrightarrow> ?J i \<in>etran" in allE, erule impE, assumption)
+ apply(erule_tac x=l and P="\<lambda>i. H i \<longrightarrow> J i \<in>etran" for H J in allE, erule impE, assumption)
apply(erule etranE,simp)
apply(erule_tac x=l in allE, erule impE, assumption,
- erule_tac x=1 and P="\<lambda>j. (?H j) \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
+ erule_tac x=1 and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE,simp)
apply(rule nth_tl_if)
- apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
apply simp
apply(rule tl_zero,force)
apply force
- apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
apply(rule disjI2)
apply simp
apply clarify
apply(case_tac j,simp)
apply(rule tl_zero)
- apply(erule_tac x=i and P="\<lambda>i. ?H i \<longrightarrow> ?J i \<in>etran" in allE, erule impE, assumption)
- apply(erule_tac x=i and P="\<lambda>i. ?H i \<longrightarrow> ?J i \<in>etran" in allE, erule impE, assumption)
+ apply(erule_tac x=i and P="\<lambda>i. H i \<longrightarrow> J i \<in>etran" for H J in allE, erule impE, assumption)
+ apply(erule_tac x=i and P="\<lambda>i. H i \<longrightarrow> J i \<in>etran" for H J in allE, erule impE, assumption)
apply(force elim:etranE intro:Env)
apply force
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
apply simp
apply(rule tl_zero)
apply(rule tl_zero,force)
apply force
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
apply force
-apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
done
lemma aux_onlyif [rule_format]: "\<forall>xs s. (xs, s)#ys \<in> par_cptn \<longrightarrow>
@@ -769,7 +769,7 @@
apply clarify
apply(case_tac "i=ia",simp)
apply(erule CptnComp)
- apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> (?I j \<in> cptn)" in allE,simp)
+ apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> (I j \<in> cptn)" for H I in allE,simp)
apply simp
apply(erule_tac x=ia in allE)
apply(rule CptnEnv,simp)
@@ -790,7 +790,7 @@
apply(rule nth_equalityI,simp,simp)
apply simp
apply(rule nth_equalityI,simp,simp)
- apply(erule_tac x=nat and P="\<lambda>j. ?H j \<longrightarrow> (fst (?a j))=((?b j))" in allE)
+ apply(erule_tac x=nat and P="\<lambda>j. H j \<longrightarrow> (fst (a j))=((b j))" for H a b in allE)
apply(case_tac nat)
apply clarify
apply(case_tac "i=ia",simp,simp)
@@ -806,7 +806,7 @@
apply clarify
apply(rule Env)
apply simp
-apply(erule_tac x=nat and P="\<lambda>j. ?H j \<longrightarrow> (?P j \<or> ?Q j)" in allE,simp)
+apply(erule_tac x=nat and P="\<lambda>j. H j \<longrightarrow> (P j \<or> Q j)" for H P Q in allE,simp)
apply(erule disjE)
apply clarify
apply(rule_tac x=ia in exI,simp)
@@ -819,7 +819,7 @@
apply simp
apply(erule_tac x=l in allE,erule impE,assumption,erule impE, assumption,simp)
apply clarify
-apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> (?P j)\<in>etran" in allE, erule impE, assumption)
+apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> (P j)\<in>etran" for H P in allE, erule impE, assumption)
apply(case_tac "i=ia",simp,simp)
done
@@ -851,28 +851,28 @@
apply(erule cptn.cases,force,force,force)
apply(simp add:par_cp_def conjoin_def same_length_def same_program_def same_state_def compat_label_def)
apply clarify
- apply(erule_tac x=0 and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in all_dupE)
+ apply(erule_tac x=0 and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in all_dupE)
apply(subgoal_tac "a = xs")
apply(subgoal_tac "b = s",simp)
prefer 3
- apply(erule_tac x=0 and P="\<lambda>j. ?H j \<longrightarrow> (fst (?s j))=((?t j))" in allE)
+ apply(erule_tac x=0 and P="\<lambda>j. H j \<longrightarrow> (fst (s j))=((t j))" for H s t in allE)
apply (simp add:cp_def)
apply(rule nth_equalityI,simp,simp)
prefer 2
apply(erule_tac x=0 in allE)
apply (simp add:cp_def)
- apply(erule_tac x=0 and P="\<lambda>j. ?H j \<longrightarrow> (\<forall>i. ?T i \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in allE,simp)
- apply(erule_tac x=0 and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
+ apply(erule_tac x=0 and P="\<lambda>j. H j \<longrightarrow> (\<forall>i. T i \<longrightarrow> (snd (d j i))=(snd (e j i)))" for H T d e in allE,simp)
+ apply(erule_tac x=0 and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE,simp)
apply(erule_tac x=list in allE)
apply(rule_tac x="map tl clist" in exI,simp)
apply(rule conjI)
apply clarify
apply(case_tac j,simp)
apply(erule_tac x=i in allE, erule impE, assumption,
- erule_tac x="0" and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
+ erule_tac x="0" and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE,simp)
apply(erule_tac x=i in allE, erule impE, assumption,
- erule_tac x="Suc nat" and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
+ erule_tac x="Suc nat" and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)
apply(case_tac "clist!i",simp,simp)
apply(rule conjI)
apply clarify
@@ -883,9 +883,9 @@
apply(simp add:cp_def)
apply clarify
apply simp
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)
apply(case_tac "clist!i",simp,simp)
- apply(thin_tac "?H = (\<exists>i. ?J i)")
+ apply(thin_tac "H = (\<exists>i. J i)" for H J)
apply(rule conjI)
apply clarify
apply(erule_tac x=j in allE,erule impE, assumption,erule disjE)
@@ -895,34 +895,34 @@
apply(rule conjI)
apply(erule_tac x=i in allE)
apply(simp add:cp_def)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)
apply(case_tac "clist!i",simp,simp)
apply clarify
apply(erule_tac x=l in allE)
- apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
+ apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE)
apply clarify
apply(simp add:cp_def)
- apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
+ apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)
apply(case_tac "clist!l",simp,simp)
apply simp
apply(rule conjI)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)
apply(case_tac "clist!i",simp,simp)
apply clarify
- apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
- apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
+ apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE)
+ apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)
apply(case_tac "clist!l",simp,simp)
apply clarify
apply(erule_tac x=i in allE)
apply(simp add:cp_def)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)
apply(case_tac "clist!i",simp)
apply(rule nth_tl_if,simp,simp)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (?P j)\<in>etran" in allE, erule impE, assumption,simp)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (P j)\<in>etran" for H P in allE, erule impE, assumption,simp)
apply(simp add:cp_def)
apply clarify
apply(rule nth_tl_if)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)
apply(case_tac "clist!i",simp,simp)
apply force
apply force
@@ -940,14 +940,14 @@
apply clarify
apply(unfold same_length_def)
apply clarify
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,simp)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,simp)
apply(rule conjI)
apply(simp add:same_state_def)
apply clarify
apply(erule_tac x=i in allE, erule impE, assumption,
- erule_tac x=j and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
+ erule_tac x=j and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE)
apply(case_tac j,simp)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)
apply(case_tac "clist!i",simp,simp)
apply(rule conjI)
apply(simp add:same_program_def)
@@ -955,7 +955,7 @@
apply(rule nth_equalityI,simp,simp)
apply(case_tac j,simp)
apply clarify
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)
apply(case_tac "clist!i",simp,simp)
apply clarify
apply(simp add:compat_label_def)
@@ -967,13 +967,13 @@
apply(rule conjI)
apply(erule_tac x=i in allE)
apply(case_tac j,simp)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)
apply(case_tac "clist!i",simp,simp)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)
apply(case_tac "clist!i",simp,simp)
apply clarify
- apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
- apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
+ apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE)
+ apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)
apply(case_tac "clist!l",simp,simp)
apply(erule_tac x=l in allE,simp)
apply(rule disjI2)
@@ -982,9 +982,9 @@
apply(case_tac j,simp,simp)
apply(rule tl_zero,force)
apply force
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
apply force
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
apply clarify
apply(erule_tac x=i in allE)
apply(simp add:cp_def)
--- a/src/HOL/IMPP/EvenOdd.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/IMPP/EvenOdd.thy Mon Mar 30 11:36:30 2015 +0200
@@ -91,7 +91,7 @@
apply (rule hoare_derivs.Comp)
apply (rule_tac [2] hoare_derivs.Ass)
apply clarsimp
-apply (rule_tac Q = "%Z s. ?P Z s & Res_ok Z s" in hoare_derivs.Comp)
+apply (rule_tac Q = "%Z s. P Z s & Res_ok Z s" and P = P for P in hoare_derivs.Comp)
apply (rule export_s)
apply (rule_tac I1 = "%Z l. Z = l Arg & 0 < Z" and Q1 = "Res_ok" in Call_invariant [THEN conseq12])
apply (rule single_asm [THEN conseq2])
--- a/src/HOL/IMPP/Hoare.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/IMPP/Hoare.thy Mon Mar 30 11:36:30 2015 +0200
@@ -286,7 +286,7 @@
apply (blast) (* cut *)
apply (blast) (* weaken *)
apply (tactic {* ALLGOALS (EVERY'
- [REPEAT o Rule_Insts.thin_tac @{context} "hoare_derivs ?x ?y" [],
+ [REPEAT o Rule_Insts.thin_tac @{context} "hoare_derivs _ _" [],
simp_tac @{context}, clarify_tac @{context}, REPEAT o smp_tac @{context} 1]) *})
apply (simp_all (no_asm_use) add: triple_valid_def2)
apply (intro strip, tactic "smp_tac @{context} 2 1", blast) (* conseq *)
@@ -315,7 +315,7 @@
apply (unfold MGT_def)
apply (erule conseq12)
apply auto
-apply (case_tac "? t. <c,?s> -c-> t")
+apply (case_tac "\<exists>t. <c,s> -c-> t" for s)
apply (fast elim: com_det)
apply clarsimp
apply (drule single_stateE)
--- a/src/HOL/IMPP/Natural.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/IMPP/Natural.thy Mon Mar 30 11:36:30 2015 +0200
@@ -105,7 +105,7 @@
(* evaluation of com is deterministic *)
lemma com_det [rule_format (no_asm)]: "<c,s> -c-> t ==> (!u. <c,s> -c-> u --> u=t)"
apply (erule evalc.induct)
-apply (erule_tac [8] V = "<?c,s1> -c-> s2" in thin_rl)
+apply (erule_tac [8] V = "<c,s1> -c-> s2" for c in thin_rl)
apply (blast elim: evalc_WHILE_case)+
done
--- a/src/HOL/Induct/Com.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Induct/Com.thy Mon Mar 30 11:36:30 2015 +0200
@@ -184,7 +184,7 @@
apply blast
apply blast
apply (blast elim: exec_WHILE_case)
-apply (erule_tac V = "(?c,s2) -[?ev]-> s3" in thin_rl)
+apply (erule_tac V = "(c,s2) -[ev]-> s3" for c ev in thin_rl)
apply clarify
apply (erule exec_WHILE_case, blast+)
done
--- a/src/HOL/Library/Formal_Power_Series.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Library/Formal_Power_Series.thy Mon Mar 30 11:36:30 2015 +0200
@@ -207,8 +207,8 @@
instance fps :: (cancel_ab_semigroup_add) cancel_ab_semigroup_add
proof
fix a b c :: "'a fps"
- assume "a + b = a + c"
- then show "b = c" by (simp add: expand_fps_eq)
+ show "a + b - a = b" by (simp add: expand_fps_eq)
+ show "a - b - c = a - (b + c)" by (simp add: expand_fps_eq diff_diff_eq)
qed
instance fps :: (cancel_comm_monoid_add) cancel_comm_monoid_add ..
--- a/src/HOL/Library/Function_Algebras.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Library/Function_Algebras.thy Mon Mar 30 11:36:30 2015 +0200
@@ -71,7 +71,7 @@
by default (simp add: fun_eq_iff add.commute)
instance "fun" :: (type, cancel_ab_semigroup_add) cancel_ab_semigroup_add
- by default simp
+ by default (simp_all add: fun_eq_iff diff_diff_eq)
instance "fun" :: (type, monoid_add) monoid_add
by default (simp_all add: fun_eq_iff)
--- a/src/HOL/Library/Library.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Library/Library.thy Mon Mar 30 11:36:30 2015 +0200
@@ -42,7 +42,7 @@
Mapping
Monad_Syntax
More_List
- Multiset
+ Multiset_Order
Numeral_Type
NthRoot_Limits
OptionalSugar
--- a/src/HOL/Library/Multiset.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Library/Multiset.thy Mon Mar 30 11:36:30 2015 +0200
@@ -1,6 +1,9 @@
(* Title: HOL/Library/Multiset.thy
Author: Tobias Nipkow, Markus Wenzel, Lawrence C Paulson, Norbert Voelker
Author: Andrei Popescu, TU Muenchen
+ Author: Jasmin Blanchette, Inria, LORIA, MPII
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Mathias Fleury, MPII
*)
section {* (Finite) multisets *}
@@ -92,8 +95,11 @@
lift_definition plus_multiset :: "'a multiset => 'a multiset => 'a multiset" is "\<lambda>M N. (\<lambda>a. M a + N a)"
by (rule union_preserves_multiset)
+lift_definition minus_multiset :: "'a multiset => 'a multiset => 'a multiset" is "\<lambda> M N. \<lambda>a. M a - N a"
+by (rule diff_preserves_multiset)
+
instance
-by default (transfer, simp add: fun_eq_iff)+
+ by default (transfer, simp add: fun_eq_iff)+
end
@@ -126,9 +132,6 @@
instantiation multiset :: (type) comm_monoid_diff
begin
-lift_definition minus_multiset :: "'a multiset => 'a multiset => 'a multiset" is "\<lambda> M N. \<lambda>a. M a - N a"
-by (rule diff_preserves_multiset)
-
instance
by default (transfer, simp add: fun_eq_iff)+
@@ -382,8 +385,7 @@
lemma multi_psub_self[simp]: "(A::'a multiset) < A = False"
by simp
-lemma mset_less_add_bothsides:
- "T + {#x#} < S + {#x#} \<Longrightarrow> T < S"
+lemma mset_less_add_bothsides: "N + {#x#} < M + {#x#} \<Longrightarrow> N < M"
by (fact add_less_imp_less_right)
lemma mset_less_empty_nonempty:
@@ -593,6 +595,10 @@
lemma set_of_mono: "A \<le> B \<Longrightarrow> set_of A \<subseteq> set_of B"
by (metis mset_leD subsetI mem_set_of_iff)
+lemma ball_set_of_iff: "(\<forall>x \<in> set_of M. P x) \<longleftrightarrow> (\<forall>x. x \<in># M \<longrightarrow> P x)"
+ by auto
+
+
subsubsection {* Size *}
definition wcount where "wcount f M = (\<lambda>x. count M x * Suc (f x))"
@@ -895,12 +901,24 @@
translations
"{#e. x:#M#}" == "CONST image_mset (%x. e) M"
+syntax (xsymbols)
+ "_comprehension2_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset"
+ ("({#_/. _ \<in># _#})")
+translations
+ "{#e. x \<in># M#}" == "CONST image_mset (\<lambda>x. e) M"
+
syntax
- "_comprehension2_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"
+ "_comprehension3_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"
("({#_/ | _ :# _./ _#})")
translations
"{#e | x:#M. P#}" => "{#e. x :# {# x:#M. P#}#}"
+syntax
+ "_comprehension4_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"
+ ("({#_/ | _ \<in># _./ _#})")
+translations
+ "{#e | x\<in>#M. P#}" => "{#e. x \<in># {# x\<in>#M. P#}#}"
+
text {*
This allows to write not just filters like @{term "{#x:#M. x<c#}"}
but also images like @{term "{#x+x. x:#M #}"} and @{term [source]
@@ -908,6 +926,9 @@
@{term "{#x+x|x:#M. x<c#}"}.
*}
+lemma in_image_mset: "y \<in># {#f x. x \<in># M#} \<longleftrightarrow> y \<in> f ` set_of M"
+ by (metis mem_set_of_iff set_of_image_mset)
+
functor image_mset: image_mset
proof -
fix f g show "image_mset f \<circ> image_mset g = image_mset (f \<circ> g)"
@@ -924,7 +945,19 @@
qed
qed
-declare image_mset.identity [simp]
+declare
+ image_mset.id [simp]
+ image_mset.identity [simp]
+
+lemma image_mset_id[simp]: "image_mset id x = x"
+ unfolding id_def by auto
+
+lemma image_mset_cong: "(\<And>x. x \<in># M \<Longrightarrow> f x = g x) \<Longrightarrow> {#f x. x \<in># M#} = {#g x. x \<in># M#}"
+ by (induct M) auto
+
+lemma image_mset_cong_pair:
+ "(\<forall>x y. (x, y) \<in># M \<longrightarrow> f x y = g x y) \<Longrightarrow> {#f x y. (x, y) \<in># M#} = {#g x y. (x, y) \<in># M#}"
+ by (metis image_mset_cong split_cong)
subsection {* Further conversions *}
@@ -942,7 +975,7 @@
by (induct xs) simp_all
lemma multiset_of_zero_iff[simp]: "(multiset_of x = {#}) = (x = [])"
-by (induct x) auto
+ by (induct x) auto
lemma multiset_of_zero_iff_right[simp]: "({#} = multiset_of x) = (x = [])"
by (induct x) auto
@@ -1059,10 +1092,6 @@
"multiset_of (insort x xs) = multiset_of xs + {#x#}"
by (induct xs) (simp_all add: ac_simps)
-lemma in_multiset_of:
- "x \<in># multiset_of xs \<longleftrightarrow> x \<in> set xs"
- by (induct xs) simp_all
-
lemma multiset_of_map:
"multiset_of (map f xs) = image_mset f (multiset_of xs)"
by (induct xs) simp_all
@@ -1098,6 +1127,9 @@
by (auto elim!: Set.set_insert)
qed -- {* TODO: maybe define @{const multiset_of_set} also in terms of @{const Abs_multiset} *}
+lemma elem_multiset_of_set[simp, intro]: "finite A \<Longrightarrow> x \<in># multiset_of_set A \<longleftrightarrow> x \<in> A"
+ by (induct A rule: finite_induct) simp_all
+
context linorder
begin
@@ -1186,6 +1218,14 @@
end
+lemma comp_fun_commute_plus_mset[simp]: "comp_fun_commute (op + \<Colon> 'a multiset \<Rightarrow> _ \<Rightarrow> _)"
+ by default (simp add: add_ac comp_def)
+
+declare comp_fun_commute.fold_mset_insert[OF comp_fun_commute_plus_mset, simp]
+
+lemma in_mset_fold_plus_iff[iff]: "x \<in># Multiset.fold (op +) M NN \<longleftrightarrow> x \<in># M \<or> (\<exists>N. N \<in># NN \<and> x \<in># N)"
+ by (induct NN) auto
+
notation times (infixl "*" 70)
notation Groups.one ("1")
@@ -1210,6 +1250,22 @@
end
+lemma msetsum_diff:
+ fixes M N :: "('a \<Colon> ordered_cancel_comm_monoid_diff) multiset"
+ shows "N \<le> M \<Longrightarrow> msetsum (M - N) = msetsum M - msetsum N"
+ by (metis add_diff_cancel_left' msetsum.union ordered_cancel_comm_monoid_diff_class.add_diff_inverse)
+
+abbreviation Union_mset :: "'a multiset multiset \<Rightarrow> 'a multiset" where
+ "Union_mset MM \<equiv> msetsum MM"
+
+notation (xsymbols) Union_mset ("\<Union>#_" [900] 900)
+
+lemma set_of_Union_mset[simp]: "set_of (\<Union># MM) = (\<Union>M \<in> set_of MM. set_of M)"
+ by (induct MM) auto
+
+lemma in_Union_mset_iff[iff]: "x \<in># \<Union># MM \<longleftrightarrow> (\<exists>M. M \<in># MM \<and> x \<in># M)"
+ by (induct MM) auto
+
syntax
"_msetsum_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add"
("(3SUM _:#_. _)" [0, 51, 10] 10)
@@ -1338,6 +1394,46 @@
lemma mcard_filter_lesseq[simp]: "mcard (Multiset.filter f M) \<le> mcard M"
by (rule mcard_mono[OF multiset_filter_subset])
+lemma mcard_1_singleton:
+ assumes card: "mcard AA = 1"
+ shows "\<exists>A. AA = {#A#}"
+ using card by (cases AA) auto
+
+lemma mcard_Diff_subset:
+ assumes "M \<le> M'"
+ shows "mcard (M' - M) = mcard M' - mcard M"
+ by (metis add_diff_cancel_left' assms mcard_plus mset_le_exists_conv)
+
+
+subsection {* Replicate operation *}
+
+definition replicate_mset :: "nat \<Rightarrow> 'a \<Rightarrow> 'a multiset" where
+ "replicate_mset n x = ((op + {#x#}) ^^ n) {#}"
+
+lemma replicate_mset_0[simp]: "replicate_mset 0 x = {#}"
+ unfolding replicate_mset_def by simp
+
+lemma replicate_mset_Suc[simp]: "replicate_mset (Suc n) x = replicate_mset n x + {#x#}"
+ unfolding replicate_mset_def by (induct n) (auto intro: add.commute)
+
+lemma in_replicate_mset[simp]: "x \<in># replicate_mset n y \<longleftrightarrow> n > 0 \<and> x = y"
+ unfolding replicate_mset_def by (induct n) simp_all
+
+lemma count_replicate_mset[simp]: "count (replicate_mset n x) y = (if y = x then n else 0)"
+ unfolding replicate_mset_def by (induct n) simp_all
+
+lemma set_of_replicate_mset_subset[simp]: "set_of (replicate_mset n x) = (if n = 0 then {} else {x})"
+ by (auto split: if_splits)
+
+lemma mcard_replicate_mset[simp]: "mcard (replicate_mset n M) = n"
+ by (induct n, simp_all)
+
+lemma count_le_replicate_mset_le: "n \<le> count M x \<longleftrightarrow> replicate_mset n x \<le> M"
+ by (auto simp add: assms mset_less_eqI) (metis count_replicate_mset less_eq_multiset.rep_eq)
+
+lemma filter_eq_replicate_mset: "{#y \<in># D. y = x#} = replicate_mset (count D x) x"
+ by (induct D) simp_all
+
subsection {* Alternative representations *}
@@ -1658,7 +1754,7 @@
apply (simp (no_asm))
apply (rule_tac x = "(K - {#a#}) + Ka" in exI)
apply (simp (no_asm_simp) add: add.assoc [symmetric])
- apply (drule_tac f = "\<lambda>M. M - {#a#}" and x="?S + ?T" in arg_cong)
+ apply (drule_tac f = "\<lambda>M. M - {#a#}" and x="S + T" for S T in arg_cong)
apply (simp add: diff_union_single_conv)
apply (simp (no_asm_use) add: trans_def)
apply blast
@@ -1669,7 +1765,7 @@
apply (rule conjI)
apply (simp add: multiset_eq_iff split: nat_diff_split)
apply (rule conjI)
- apply (drule_tac f = "\<lambda>M. M - {#a#}" and x="?S + ?T" in arg_cong, simp)
+ apply (drule_tac f = "\<lambda>M. M - {#a#}" and x="S + T" for S T in arg_cong, simp)
apply (simp add: multiset_eq_iff split: nat_diff_split)
apply (simp (no_asm_use) add: trans_def)
apply blast
@@ -1692,7 +1788,7 @@
apply (simp add: mult1_def set_of_def, blast)
txt {* Now we know @{term "J' \<noteq> {#}"}. *}
apply (cut_tac M = K and P = "\<lambda>x. (x, a) \<in> r" in multiset_partition)
-apply (erule_tac P = "\<forall>k \<in> set_of K. ?P k" in rev_mp)
+apply (erule_tac P = "\<forall>k \<in> set_of K. P k" for P in rev_mp)
apply (erule ssubst)
apply (simp add: Ball_def, auto)
apply (subgoal_tac
@@ -1790,7 +1886,7 @@
qed (auto simp add: le_multiset_def intro: union_less_mono2)
-subsection {* Termination proofs with multiset orders *}
+subsubsection {* Termination proofs with multiset orders *}
lemma multi_member_skip: "x \<in># XS \<Longrightarrow> x \<in># {# y #} + XS"
and multi_member_this: "x \<in># {# x #} + XS"
@@ -2082,9 +2178,7 @@
then show ?thesis by simp
qed
-lemma [code_unfold]:
- "x \<in># multiset_of xs \<longleftrightarrow> x \<in> set xs"
- by (simp add: in_multiset_of)
+declare in_multiset_in_set [code_unfold]
lemma [code]:
"count (multiset_of xs) x = fold (\<lambda>y. if x = y then Suc else id) xs 0"
@@ -2094,25 +2188,19 @@
then show ?thesis by simp
qed
-lemma [code]:
- "set_of (multiset_of xs) = set xs"
- by simp
-
-lemma [code]:
- "sorted_list_of_multiset (multiset_of xs) = sort xs"
- by (induct xs) simp_all
+declare set_of_multiset_of [code]
+
+declare sorted_list_of_multiset_multiset_of [code]
lemma [code]: -- {* not very efficient, but representation-ignorant! *}
"multiset_of_set A = multiset_of (sorted_list_of_set A)"
apply (cases "finite A")
apply simp_all
apply (induct A rule: finite_induct)
- apply (simp_all add: union_commute)
+ apply (simp_all add: add.commute)
done
-lemma [code]:
- "mcard (multiset_of xs) = length xs"
- by (simp add: mcard_multiset_of)
+declare mcard_multiset_of [code]
fun ms_lesseq_impl :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool option" where
"ms_lesseq_impl [] ys = Some (ys \<noteq> [])"
@@ -2287,7 +2375,7 @@
have "multiset_of xs' = {#x#} + multiset_of xsa"
unfolding xsa_def using j_len nth_j
by (metis (no_types) ab_semigroup_add_class.add_ac(1) append_take_drop_id Cons_nth_drop_Suc
- multiset_of.simps(2) union_code union_commute)
+ multiset_of.simps(2) union_code add.commute)
hence ms_x: "multiset_of xsa = multiset_of xs"
by (metis Cons.prems add.commute add_right_imp_eq multiset_of.simps(2))
then obtain ysa where
@@ -2344,7 +2432,7 @@
by (rule image_mset.id)
next
show "\<And>f g. image_mset (g \<circ> f) = image_mset g \<circ> image_mset f"
- unfolding comp_def by (rule ext) (simp add: image_mset.compositionality comp_def)
+ unfolding comp_def by (rule ext) (simp add: comp_def image_mset.compositionality)
next
fix X :: "'a multiset"
show "\<And>f g. (\<And>z. z \<in> set_of X \<Longrightarrow> f z = g z) \<Longrightarrow> image_mset f X = image_mset g X"
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Multiset_Order.thy Mon Mar 30 11:36:30 2015 +0200
@@ -0,0 +1,308 @@
+(* Title: HOL/Library/Multiset_Order.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Jasmin Blanchette, Inria, LORIA, MPII
+*)
+
+section {* More Theorems about the Multiset Order *}
+
+theory Multiset_Order
+imports Multiset
+begin
+
+subsubsection {* Alternative characterizations *}
+
+context order
+begin
+
+lemma reflp_le: "reflp (op \<le>)"
+ unfolding reflp_def by simp
+
+lemma antisymP_le: "antisymP (op \<le>)"
+ unfolding antisym_def by auto
+
+lemma transp_le: "transp (op \<le>)"
+ unfolding transp_def by auto
+
+lemma irreflp_less: "irreflp (op <)"
+ unfolding irreflp_def by simp
+
+lemma antisymP_less: "antisymP (op <)"
+ unfolding antisym_def by auto
+
+lemma transp_less: "transp (op <)"
+ unfolding transp_def by auto
+
+lemmas le_trans = transp_le[unfolded transp_def, rule_format]
+
+lemma order_mult: "class.order
+ (\<lambda>M N. (M, N) \<in> mult {(x, y). x < y} \<or> M = N)
+ (\<lambda>M N. (M, N) \<in> mult {(x, y). x < y})"
+ (is "class.order ?le ?less")
+proof -
+ have irrefl: "\<And>M :: 'a multiset. \<not> ?less M M"
+ proof
+ fix M :: "'a multiset"
+ have "trans {(x'::'a, x). x' < x}"
+ by (rule transI) simp
+ moreover
+ assume "(M, M) \<in> mult {(x, y). x < y}"
+ ultimately have "\<exists>I J K. M = I + J \<and> M = I + K
+ \<and> J \<noteq> {#} \<and> (\<forall>k\<in>set_of K. \<exists>j\<in>set_of J. (k, j) \<in> {(x, y). x < y})"
+ by (rule mult_implies_one_step)
+ then obtain I J K where "M = I + J" and "M = I + K"
+ and "J \<noteq> {#}" and "(\<forall>k\<in>set_of K. \<exists>j\<in>set_of J. (k, j) \<in> {(x, y). x < y})" by blast
+ then have aux1: "K \<noteq> {#}" and aux2: "\<forall>k\<in>set_of K. \<exists>j\<in>set_of K. k < j" by auto
+ have "finite (set_of K)" by simp
+ moreover note aux2
+ ultimately have "set_of K = {}"
+ by (induct rule: finite_induct)
+ (simp, metis (mono_tags) insert_absorb insert_iff insert_not_empty less_irrefl less_trans)
+ with aux1 show False by simp
+ qed
+ have trans: "\<And>K M N :: 'a multiset. ?less K M \<Longrightarrow> ?less M N \<Longrightarrow> ?less K N"
+ unfolding mult_def by (blast intro: trancl_trans)
+ show "class.order ?le ?less"
+ by default (auto simp add: le_multiset_def irrefl dest: trans)
+qed
+
+text {* The Dershowitz--Manna ordering: *}
+
+definition less_multiset\<^sub>D\<^sub>M where
+ "less_multiset\<^sub>D\<^sub>M M N \<longleftrightarrow>
+ (\<exists>X Y. X \<noteq> {#} \<and> X \<le> N \<and> M = (N - X) + Y \<and> (\<forall>k. k \<in># Y \<longrightarrow> (\<exists>a. a \<in># X \<and> k < a)))"
+
+
+text {* The Huet--Oppen ordering: *}
+
+definition less_multiset\<^sub>H\<^sub>O where
+ "less_multiset\<^sub>H\<^sub>O M N \<longleftrightarrow> M \<noteq> N \<and> (\<forall>y. count N y < count M y \<longrightarrow> (\<exists>x. y < x \<and> count M x < count N x))"
+
+lemma mult_imp_less_multiset\<^sub>H\<^sub>O: "(M, N) \<in> mult {(x, y). x < y} \<Longrightarrow> less_multiset\<^sub>H\<^sub>O M N"
+proof (unfold mult_def less_multiset\<^sub>H\<^sub>O_def, induct rule: trancl_induct)
+ case (base P)
+ then show ?case unfolding mult1_def by force
+next
+ case (step N P)
+ from step(2) obtain M0 a K where
+ *: "P = M0 + {#a#}" "N = M0 + K" "\<And>b. b \<in># K \<Longrightarrow> b < a"
+ unfolding mult1_def by blast
+ then have count_K_a: "count K a = 0" by auto
+ with step(3) *(1,2) have "M \<noteq> P" by (force dest: *(3) split: if_splits)
+ moreover
+ { assume "count P a \<le> count M a"
+ with count_K_a have "count N a < count M a" unfolding *(1,2) by auto
+ with step(3) obtain z where z: "z > a" "count M z < count N z" by blast
+ with * have "count N z \<le> count P z" by force
+ with z have "\<exists>z > a. count M z < count P z" by auto
+ } note count_a = this
+ { fix y
+ assume count_y: "count P y < count M y"
+ have "\<exists>x>y. count M x < count P x"
+ proof (cases "y = a")
+ case True
+ with count_y count_a show ?thesis by auto
+ next
+ case False
+ show ?thesis
+ proof (cases "y \<in># K")
+ case True
+ with *(3) have "y < a" by simp
+ then show ?thesis by (cases "count P a \<le> count M a") (auto dest: count_a intro: less_trans)
+ next
+ case False
+ with `y \<noteq> a` have "count P y = count N y" unfolding *(1,2) by simp
+ with count_y step(3) obtain z where z: "z > y" "count M z < count N z" by auto
+ show ?thesis
+ proof (cases "z \<in># K")
+ case True
+ with *(3) have "z < a" by simp
+ with z(1) show ?thesis
+ by (cases "count P a \<le> count M a") (auto dest!: count_a intro: less_trans)
+ next
+ case False
+ with count_K_a have "count N z \<le> count P z" unfolding * by auto
+ with z show ?thesis by auto
+ qed
+ qed
+ qed
+ }
+ ultimately show ?case by blast
+qed
+
+lemma less_multiset\<^sub>D\<^sub>M_imp_mult:
+ "less_multiset\<^sub>D\<^sub>M M N \<Longrightarrow> (M, N) \<in> mult {(x, y). x < y}"
+proof -
+ assume "less_multiset\<^sub>D\<^sub>M M N"
+ then obtain X Y where
+ "X \<noteq> {#}" and "X \<le> N" and "M = N - X + Y" and "\<forall>k. k \<in># Y \<longrightarrow> (\<exists>a. a \<in># X \<and> k < a)"
+ unfolding less_multiset\<^sub>D\<^sub>M_def by blast
+ then have "(N - X + Y, N - X + X) \<in> mult {(x, y). x < y}"
+ by (intro one_step_implies_mult) (auto simp: Bex_def trans_def)
+ with `M = N - X + Y` `X \<le> N` show "(M, N) \<in> mult {(x, y). x < y}"
+ by (metis ordered_cancel_comm_monoid_diff_class.diff_add)
+qed
+
+lemma less_multiset\<^sub>H\<^sub>O_imp_less_multiset\<^sub>D\<^sub>M: "less_multiset\<^sub>H\<^sub>O M N \<Longrightarrow> less_multiset\<^sub>D\<^sub>M M N"
+unfolding less_multiset\<^sub>D\<^sub>M_def
+proof (intro iffI exI conjI)
+ assume "less_multiset\<^sub>H\<^sub>O M N"
+ then obtain z where z: "count M z < count N z"
+ unfolding less_multiset\<^sub>H\<^sub>O_def by (auto simp: multiset_eq_iff nat_neq_iff)
+ def X \<equiv> "N - M"
+ def Y \<equiv> "M - N"
+ from z show "X \<noteq> {#}" unfolding X_def by (auto simp: multiset_eq_iff not_less_eq_eq Suc_le_eq)
+ from z show "X \<le> N" unfolding X_def by auto
+ show "M = (N - X) + Y" unfolding X_def Y_def multiset_eq_iff count_union count_diff by force
+ show "\<forall>k. k \<in># Y \<longrightarrow> (\<exists>a. a \<in># X \<and> k < a)"
+ proof (intro allI impI)
+ fix k
+ assume "k \<in># Y"
+ then have "count N k < count M k" unfolding Y_def by auto
+ with `less_multiset\<^sub>H\<^sub>O M N` obtain a where "k < a" and "count M a < count N a"
+ unfolding less_multiset\<^sub>H\<^sub>O_def by blast
+ then show "\<exists>a. a \<in># X \<and> k < a" unfolding X_def by auto
+ qed
+qed
+
+lemma mult_less_multiset\<^sub>D\<^sub>M: "(M, N) \<in> mult {(x, y). x < y} \<longleftrightarrow> less_multiset\<^sub>D\<^sub>M M N"
+ by (metis less_multiset\<^sub>D\<^sub>M_imp_mult less_multiset\<^sub>H\<^sub>O_imp_less_multiset\<^sub>D\<^sub>M mult_imp_less_multiset\<^sub>H\<^sub>O)
+
+lemma mult_less_multiset\<^sub>H\<^sub>O: "(M, N) \<in> mult {(x, y). x < y} \<longleftrightarrow> less_multiset\<^sub>H\<^sub>O M N"
+ by (metis less_multiset\<^sub>D\<^sub>M_imp_mult less_multiset\<^sub>H\<^sub>O_imp_less_multiset\<^sub>D\<^sub>M mult_imp_less_multiset\<^sub>H\<^sub>O)
+
+lemmas mult\<^sub>D\<^sub>M = mult_less_multiset\<^sub>D\<^sub>M[unfolded less_multiset\<^sub>D\<^sub>M_def]
+lemmas mult\<^sub>H\<^sub>O = mult_less_multiset\<^sub>H\<^sub>O[unfolded less_multiset\<^sub>H\<^sub>O_def]
+
+end
+
+context linorder
+begin
+
+lemma total_le: "total {(a \<Colon> 'a, b). a \<le> b}"
+ unfolding total_on_def by auto
+
+lemma total_less: "total {(a \<Colon> 'a, b). a < b}"
+ unfolding total_on_def by auto
+
+lemma linorder_mult: "class.linorder
+ (\<lambda>M N. (M, N) \<in> mult {(x, y). x < y} \<or> M = N)
+ (\<lambda>M N. (M, N) \<in> mult {(x, y). x < y})"
+proof -
+ interpret o: order
+ "(\<lambda>M N. (M, N) \<in> mult {(x, y). x < y} \<or> M = N)"
+ "(\<lambda>M N. (M, N) \<in> mult {(x, y). x < y})"
+ by (rule order_mult)
+ show ?thesis by unfold_locales (auto 0 3 simp: mult\<^sub>H\<^sub>O not_less_iff_gr_or_eq)
+qed
+
+end
+
+lemma less_multiset_less_multiset\<^sub>H\<^sub>O:
+ "M \<subset># N \<longleftrightarrow> less_multiset\<^sub>H\<^sub>O M N"
+ unfolding less_multiset_def mult\<^sub>H\<^sub>O less_multiset\<^sub>H\<^sub>O_def ..
+
+lemmas less_multiset\<^sub>D\<^sub>M = mult\<^sub>D\<^sub>M[folded less_multiset_def]
+lemmas less_multiset\<^sub>H\<^sub>O = mult\<^sub>H\<^sub>O[folded less_multiset_def]
+
+lemma le_multiset\<^sub>H\<^sub>O:
+ fixes M N :: "('a \<Colon> linorder) multiset"
+ shows "M \<subseteq># N \<longleftrightarrow> (\<forall>y. count N y < count M y \<longrightarrow> (\<exists>x. y < x \<and> count M x < count N x))"
+ by (auto simp: le_multiset_def less_multiset\<^sub>H\<^sub>O)
+
+lemma wf_less_multiset: "wf {(M \<Colon> ('a \<Colon> wellorder) multiset, N). M \<subset># N}"
+ unfolding less_multiset_def by (auto intro: wf_mult wf)
+
+lemma order_multiset: "class.order
+ (le_multiset :: ('a \<Colon> order) multiset \<Rightarrow> ('a \<Colon> order) multiset \<Rightarrow> bool)
+ (less_multiset :: ('a \<Colon> order) multiset \<Rightarrow> ('a \<Colon> order) multiset \<Rightarrow> bool)"
+ by unfold_locales
+
+lemma linorder_multiset: "class.linorder
+ (le_multiset :: ('a \<Colon> linorder) multiset \<Rightarrow> ('a \<Colon> linorder) multiset \<Rightarrow> bool)
+ (less_multiset :: ('a \<Colon> linorder) multiset \<Rightarrow> ('a \<Colon> linorder) multiset \<Rightarrow> bool)"
+ by unfold_locales (fastforce simp add: less_multiset\<^sub>H\<^sub>O le_multiset_def not_less_iff_gr_or_eq)
+
+interpretation multiset_linorder: linorder
+ "le_multiset :: ('a \<Colon> linorder) multiset \<Rightarrow> ('a \<Colon> linorder) multiset \<Rightarrow> bool"
+ "less_multiset :: ('a \<Colon> linorder) multiset \<Rightarrow> ('a \<Colon> linorder) multiset \<Rightarrow> bool"
+ by (rule linorder_multiset)
+
+interpretation multiset_wellorder: wellorder
+ "le_multiset :: ('a \<Colon> wellorder) multiset \<Rightarrow> ('a \<Colon> wellorder) multiset \<Rightarrow> bool"
+ "less_multiset :: ('a \<Colon> wellorder) multiset \<Rightarrow> ('a \<Colon> wellorder) multiset \<Rightarrow> bool"
+ by unfold_locales (blast intro: wf_less_multiset[unfolded wf_def, rule_format])
+
+lemma le_multiset_total:
+ fixes M N :: "('a \<Colon> linorder) multiset"
+ shows "\<not> M \<subseteq># N \<Longrightarrow> N \<subseteq># M"
+ by (metis multiset_linorder.le_cases)
+
+lemma less_eq_imp_le_multiset:
+ fixes M N :: "('a \<Colon> linorder) multiset"
+ shows "M \<le> N \<Longrightarrow> M \<subseteq># N"
+ unfolding le_multiset_def less_multiset\<^sub>H\<^sub>O
+ by (auto dest: leD simp add: less_eq_multiset.rep_eq)
+
+lemma less_multiset_right_total:
+ fixes M :: "('a \<Colon> linorder) multiset"
+ shows "M \<subset># M + {#undefined#}"
+ unfolding le_multiset_def less_multiset\<^sub>H\<^sub>O by simp
+
+lemma le_multiset_empty_left[simp]:
+ fixes M :: "('a \<Colon> linorder) multiset"
+ shows "{#} \<subseteq># M"
+ by (simp add: less_eq_imp_le_multiset)
+
+lemma le_multiset_empty_right[simp]:
+ fixes M :: "('a \<Colon> linorder) multiset"
+ shows "M \<noteq> {#} \<Longrightarrow> \<not> M \<subseteq># {#}"
+ by (metis le_multiset_empty_left multiset_order.antisym)
+
+lemma less_multiset_empty_left[simp]:
+ fixes M :: "('a \<Colon> linorder) multiset"
+ shows "M \<noteq> {#} \<Longrightarrow> {#} \<subset># M"
+ by (simp add: less_multiset\<^sub>H\<^sub>O)
+
+lemma less_multiset_empty_right[simp]:
+ fixes M :: "('a \<Colon> linorder) multiset"
+ shows "\<not> M \<subset># {#}"
+ using le_empty less_multiset\<^sub>D\<^sub>M by blast
+
+lemma
+ fixes M N :: "('a \<Colon> linorder) multiset"
+ shows
+ le_multiset_plus_left[simp]: "N \<subseteq># (M + N)" and
+ le_multiset_plus_right[simp]: "M \<subseteq># (M + N)"
+ using [[metis_verbose = false]] by (metis less_eq_imp_le_multiset mset_le_add_left add.commute)+
+
+lemma
+ fixes M N :: "('a \<Colon> linorder) multiset"
+ shows
+ less_multiset_plus_plus_left_iff[simp]: "M + N \<subset># M' + N \<longleftrightarrow> M \<subset># M'" and
+ less_multiset_plus_plus_right_iff[simp]: "M + N \<subset># M + N' \<longleftrightarrow> N \<subset># N'"
+ unfolding less_multiset\<^sub>H\<^sub>O by auto
+
+lemma add_eq_self_empty_iff: "M + N = M \<longleftrightarrow> N = {#}"
+ by (metis add.commute add_diff_cancel_right' monoid_add_class.add.left_neutral)
+
+lemma
+ fixes M N :: "('a \<Colon> linorder) multiset"
+ shows
+ less_multiset_plus_left_nonempty[simp]: "M \<noteq> {#} \<Longrightarrow> N \<subset># M + N" and
+ less_multiset_plus_right_nonempty[simp]: "N \<noteq> {#} \<Longrightarrow> M \<subset># M + N"
+ using [[metis_verbose = false]]
+ by (metis add.right_neutral less_multiset_empty_left less_multiset_plus_plus_right_iff
+ add.commute)+
+
+lemma ex_gt_imp_less_multiset: "(\<exists>y \<Colon> 'a \<Colon> linorder. y \<in># N \<and> (\<forall>x. x \<in># M \<longrightarrow> x < y)) \<Longrightarrow> M \<subset># N"
+ unfolding less_multiset\<^sub>H\<^sub>O by (metis less_irrefl less_nat_zero_code not_gr0)
+
+lemma ex_gt_count_imp_less_multiset:
+ "(\<forall>y \<Colon> 'a \<Colon> linorder. y \<in># M + N \<longrightarrow> y \<le> x) \<Longrightarrow> count M x < count N x \<Longrightarrow> M \<subset># N"
+ unfolding less_multiset\<^sub>H\<^sub>O by (metis add.left_neutral add_lessD1 dual_order.strict_iff_order
+ less_not_sym mset_leD mset_le_add_left)
+
+lemma union_less_diff_plus: "P \<le> M \<Longrightarrow> N \<subset># P \<Longrightarrow> M - P + N \<subset># M"
+ by (drule ordered_cancel_comm_monoid_diff_class.diff_add[symmetric]) (metis union_less_mono2)
+
+end
--- a/src/HOL/Library/Polynomial.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Library/Polynomial.thy Mon Mar 30 11:36:30 2015 +0200
@@ -559,13 +559,32 @@
end
-instance poly :: (cancel_comm_monoid_add) cancel_comm_monoid_add
-proof
+instantiation poly :: (cancel_comm_monoid_add) cancel_comm_monoid_add
+begin
+
+lift_definition minus_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
+ is "\<lambda>p q n. coeff p n - coeff q n"
+proof (rule almost_everywhere_zeroI)
+ fix q p :: "'a poly" and i
+ assume "max (degree q) (degree p) < i"
+ then show "coeff p i - coeff q i = 0"
+ by (simp add: coeff_eq_0)
+qed
+
+lemma coeff_diff [simp]:
+ "coeff (p - q) n = coeff p n - coeff q n"
+ by (simp add: minus_poly.rep_eq)
+
+instance proof
fix p q r :: "'a poly"
- assume "p + q = p + r" thus "q = r"
+ show "p + q - p = q"
by (simp add: poly_eq_iff)
+ show "p - q - r = p - (q + r)"
+ by (simp add: poly_eq_iff diff_diff_eq)
qed
+end
+
instantiation poly :: (ab_group_add) ab_group_add
begin
@@ -578,22 +597,9 @@
by (simp add: coeff_eq_0)
qed
-lift_definition minus_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
- is "\<lambda>p q n. coeff p n - coeff q n"
-proof (rule almost_everywhere_zeroI)
- fix q p :: "'a poly" and i
- assume "max (degree q) (degree p) < i"
- then show "coeff p i - coeff q i = 0"
- by (simp add: coeff_eq_0)
-qed
-
lemma coeff_minus [simp]: "coeff (- p) n = - coeff p n"
by (simp add: uminus_poly.rep_eq)
-lemma coeff_diff [simp]:
- "coeff (p - q) n = coeff p n - coeff q n"
- by (simp add: minus_poly.rep_eq)
-
instance proof
fix p q :: "'a poly"
show "- p + p = 0"
@@ -641,20 +647,27 @@
using degree_add_eq_right [of q p]
by (simp add: add.commute)
-lemma degree_minus [simp]: "degree (- p) = degree p"
+lemma degree_minus [simp]:
+ "degree (- p) = degree p"
unfolding degree_def by simp
-lemma degree_diff_le_max: "degree (p - q) \<le> max (degree p) (degree q)"
+lemma degree_diff_le_max:
+ fixes p q :: "'a :: ab_group_add poly"
+ shows "degree (p - q) \<le> max (degree p) (degree q)"
using degree_add_le [where p=p and q="-q"]
by simp
lemma degree_diff_le:
- "\<lbrakk>degree p \<le> n; degree q \<le> n\<rbrakk> \<Longrightarrow> degree (p - q) \<le> n"
- using degree_add_le [of p n "- q"] by simp
+ fixes p q :: "'a :: ab_group_add poly"
+ assumes "degree p \<le> n" and "degree q \<le> n"
+ shows "degree (p - q) \<le> n"
+ using assms degree_add_le [of p n "- q"] by simp
lemma degree_diff_less:
- "\<lbrakk>degree p < n; degree q < n\<rbrakk> \<Longrightarrow> degree (p - q) < n"
- using degree_add_less [of p n "- q"] by simp
+ fixes p q :: "'a :: ab_group_add poly"
+ assumes "degree p < n" and "degree q < n"
+ shows "degree (p - q) < n"
+ using assms degree_add_less [of p n "- q"] by simp
lemma add_monom: "monom a n + monom b n = monom (a + b) n"
by (rule poly_eqI) simp
--- a/src/HOL/Library/Product_plus.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Library/Product_plus.thy Mon Mar 30 11:36:30 2015 +0200
@@ -98,7 +98,7 @@
instance prod ::
(cancel_ab_semigroup_add, cancel_ab_semigroup_add) cancel_ab_semigroup_add
- by default (simp add: prod_eq_iff)
+ by default (simp_all add: prod_eq_iff diff_diff_eq)
instance prod ::
(cancel_comm_monoid_add, cancel_comm_monoid_add) cancel_comm_monoid_add ..
--- a/src/HOL/MicroJava/DFA/Listn.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/MicroJava/DFA/Listn.thy Mon Mar 30 11:36:30 2015 +0200
@@ -345,7 +345,7 @@
apply (erule_tac x = "{a. \<exists>xs. size xs = k \<and> a#xs:M}" in allE)
apply (erule impE)
apply blast
-apply (thin_tac "\<exists>x xs. ?P x xs")
+apply (thin_tac "\<exists>x xs. P x xs" for P)
apply clarify
apply (rename_tac maxA xs)
apply (erule_tac x = "{ys. size ys = size xs \<and> maxA#ys \<in> M}" in allE)
--- a/src/HOL/MicroJava/J/JTypeSafe.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/MicroJava/J/JTypeSafe.thy Mon Mar 30 11:36:30 2015 +0200
@@ -137,20 +137,20 @@
apply( clarsimp)
apply( drule (3) Call_lemma)
apply( clarsimp simp add: wf_java_mdecl_def)
-apply( erule_tac V = "method ?sig ?x = ?y" in thin_rl)
+apply( erule_tac V = "method sig x = y" for sig x y in thin_rl)
apply( drule spec, erule impE, erule_tac [2] notE impE, tactic "assume_tac @{context} 2")
apply( rule conformsI)
apply( erule conforms_heapD)
apply( rule lconf_ext)
apply( force elim!: Call_lemma2)
apply( erule conf_hext, erule (1) conf_obj_AddrI)
-apply( erule_tac V = "?E\<turnstile>?blk\<surd>" in thin_rl)
+apply( erule_tac V = "E\<turnstile>blk\<surd>" for E blk in thin_rl)
apply (simp add: conforms_def)
apply( erule conjE)
apply( drule spec, erule (1) impE)
apply( drule spec, erule (1) impE)
-apply( erule_tac V = "?E\<turnstile>res::?rT" in thin_rl)
+apply( erule_tac V = "E\<turnstile>res::rT" for E rT in thin_rl)
apply( clarify)
apply( rule conjI)
apply( fast intro: hext_trans)
@@ -244,7 +244,7 @@
THEN_ALL_NEW (full_simp_tac @{context}), REPEAT o (etac conjE), hyp_subst_tac @{context}] 3*})
-- "for if"
-apply( tactic {* (Induct_Tacs.case_tac @{context} "the_Bool v" [] THEN_ALL_NEW
+apply( tactic {* (Induct_Tacs.case_tac @{context} "the_Bool v" [] NONE THEN_ALL_NEW
(asm_full_simp_tac @{context})) 7*})
apply (tactic "forward_hyp_tac @{context}")
@@ -291,7 +291,7 @@
-- "5 While"
prefer 5
-apply(erule_tac V = "?a \<longrightarrow> ?b" in thin_rl)
+apply(erule_tac V = "a \<longrightarrow> b" for a b in thin_rl)
apply(drule (1) ty_expr_ty_exprs_wt_stmt.Loop)
apply(force elim: hext_trans)
--- a/src/HOL/MicroJava/J/WellForm.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/MicroJava/J/WellForm.thy Mon Mar 30 11:36:30 2015 +0200
@@ -419,7 +419,7 @@
apply( simp)
apply( drule map_add_SomeD)
apply( erule disjE)
-apply( erule_tac V = "?P --> ?Q" in thin_rl)
+apply( erule_tac V = "P --> Q" for P Q in thin_rl)
apply (frule map_of_SomeD)
apply (clarsimp simp add: wf_cdecl_def)
apply( clarify)
@@ -454,7 +454,7 @@
apply( simp)
apply( drule map_add_SomeD)
apply( erule disjE)
-apply( erule_tac V = "?P --> ?Q" in thin_rl)
+apply( erule_tac V = "P --> Q" for P Q in thin_rl)
apply (frule map_of_SomeD)
apply (clarsimp simp add: ws_cdecl_def)
apply blast
@@ -486,7 +486,7 @@
apply( assumption)
apply( unfold map_add_def)
apply( simp (no_asm_simp) add: wf_prog_ws_prog del: split_paired_Ex)
-apply( case_tac "\<exists>z. map_of(map (\<lambda>(s,m). (s, ?C, m)) ms) sig = Some z")
+apply( case_tac "\<exists>z. map_of(map (\<lambda>(s,m). (s, C, m)) ms) sig = Some z" for C)
apply( erule exE)
apply( rotate_tac -1, frule ssubst, erule_tac [2] asm_rl)
prefer 2
--- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Mon Mar 30 11:36:30 2015 +0200
@@ -6703,7 +6703,7 @@
moreover have "?d > 0"
using as(2) by (auto simp: Suc_le_eq DIM_positive)
ultimately show "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> (\<forall>i\<in>Basis. 0 \<le> y \<bullet> i) \<and> setsum (op \<bullet> y) Basis \<le> 1"
- apply (rule_tac x="min (Min ((op \<bullet> x) ` Basis)) ?D" in exI)
+ apply (rule_tac x="min (Min ((op \<bullet> x) ` Basis)) D" for D in exI)
apply rule
defer
apply (rule, rule)
--- a/src/HOL/Multivariate_Analysis/Derivative.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Derivative.thy Mon Mar 30 11:36:30 2015 +0200
@@ -133,7 +133,7 @@
apply (simp add: LIM_zero_iff [where l = D, symmetric])
apply (simp add: Lim_within dist_norm)
apply (simp add: nonzero_norm_divide [symmetric])
- apply (simp add: 1 diff_add_eq_diff_diff ac_simps)
+ apply (simp add: 1 diff_diff_eq ac_simps)
done
qed
@@ -299,14 +299,14 @@
"(f has_derivative f') (at x within s) \<longleftrightarrow> bounded_linear f' \<and>
(\<forall>e>0. \<exists>d>0. \<forall>y\<in>s. norm(y - x) < d \<longrightarrow> norm (f y - f x - f' (y - x)) \<le> e * norm (y - x))"
unfolding has_derivative_within filterlim_def le_nhds_metric_le eventually_filtermap
- eventually_at dist_norm diff_add_eq_diff_diff
+ eventually_at dist_norm diff_diff_eq
by (force simp add: linear_0 bounded_linear.linear pos_divide_le_eq)
lemma has_derivative_within_alt2:
"(f has_derivative f') (at x within s) \<longleftrightarrow> bounded_linear f' \<and>
(\<forall>e>0. eventually (\<lambda>y. norm (f y - f x - f' (y - x)) \<le> e * norm (y - x)) (at x within s))"
unfolding has_derivative_within filterlim_def le_nhds_metric_le eventually_filtermap
- eventually_at dist_norm diff_add_eq_diff_diff
+ eventually_at dist_norm diff_diff_eq
by (force simp add: linear_0 bounded_linear.linear pos_divide_le_eq)
lemma has_derivative_at_alt:
--- a/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy Mon Mar 30 11:36:30 2015 +0200
@@ -110,7 +110,7 @@
by default (simp_all add: vec_eq_iff)
instance vec :: (cancel_ab_semigroup_add, finite) cancel_ab_semigroup_add
- by default (simp add: vec_eq_iff)
+ by default (simp_all add: vec_eq_iff diff_diff_eq)
instance vec :: (cancel_comm_monoid_add, finite) cancel_comm_monoid_add ..
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Mar 30 11:36:30 2015 +0200
@@ -5392,11 +5392,7 @@
unfolding mem_interior by auto
then have "ball (x - a) e \<subseteq> s"
unfolding subset_eq Ball_def mem_ball dist_norm
- apply auto
- apply (erule_tac x="a + xa" in allE)
- unfolding ab_group_add_class.diff_diff_eq[symmetric]
- apply auto
- done
+ by (auto simp add: diff_diff_eq)
then show "x \<in> op + a ` interior s"
unfolding image_iff
apply (rule_tac x="x - a" in bexI)
--- a/src/HOL/NSA/StarDef.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/NSA/StarDef.thy Mon Mar 30 11:36:30 2015 +0200
@@ -698,7 +698,7 @@
star_inf_def [transfer_unfold]: "inf \<equiv> *f2* inf"
instance
- by default (transfer star_inf_def, auto)+
+ by default (transfer, auto)+
end
@@ -709,7 +709,7 @@
star_sup_def [transfer_unfold]: "sup \<equiv> *f2* sup"
instance
- by default (transfer star_sup_def, auto)+
+ by default (transfer, auto)+
end
@@ -795,7 +795,7 @@
done
instance star :: (cancel_ab_semigroup_add) cancel_ab_semigroup_add
-by (intro_classes, transfer, rule add_left_imp_eq)
+by intro_classes (transfer, simp add: diff_diff_eq)+
instance star :: (cancel_comm_monoid_add) cancel_comm_monoid_add ..
@@ -850,13 +850,10 @@
declare dvd_def [transfer_refold]
-instance star :: (semiring_dvd) semiring_dvd
-apply intro_classes
-apply(transfer, rule dvd_add_times_triv_left_iff)
-apply(transfer, erule (1) dvd_addD)
-done
+instance star :: (comm_semiring_1_diff_distrib) comm_semiring_1_diff_distrib
+by intro_classes (transfer, fact right_diff_distrib')
-instance star :: (no_zero_divisors) no_zero_divisors
+instance star :: (semiring_no_zero_divisors) semiring_no_zero_divisors
by (intro_classes, transfer, rule no_zero_divisors)
instance star :: (semiring_1_cancel) semiring_1_cancel ..
@@ -865,7 +862,7 @@
instance star :: (comm_ring) comm_ring ..
instance star :: (ring_1) ring_1 ..
instance star :: (comm_ring_1) comm_ring_1 ..
-instance star :: (semiring_no_zero_divisors) semiring_no_zero_divisors ..
+instance star :: (semidom) semidom ..
instance star :: (semiring_div) semiring_div
apply intro_classes
apply(transfer, rule mod_div_equality)
@@ -1040,18 +1037,16 @@
instance star :: (semiring_numeral_div) semiring_numeral_div
apply intro_classes
-apply(transfer, erule semiring_numeral_div_class.diff_invert_add1)
-apply(transfer, erule semiring_numeral_div_class.le_add_diff_inverse2)
-apply(transfer, rule semiring_numeral_div_class.mult_div_cancel)
-apply(transfer, erule (1) semiring_numeral_div_class.div_less)
-apply(transfer, erule (1) semiring_numeral_div_class.mod_less)
-apply(transfer, erule (1) semiring_numeral_div_class.div_positive)
-apply(transfer, erule semiring_numeral_div_class.mod_less_eq_dividend)
-apply(transfer, erule semiring_numeral_div_class.pos_mod_bound)
-apply(transfer, erule semiring_numeral_div_class.pos_mod_sign)
-apply(transfer, erule semiring_numeral_div_class.mod_mult2_eq)
-apply(transfer, erule semiring_numeral_div_class.div_mult2_eq)
-apply(transfer, rule discrete)
+apply(transfer, fact semiring_numeral_div_class.le_add_diff_inverse2)
+apply(transfer, fact semiring_numeral_div_class.div_less)
+apply(transfer, fact semiring_numeral_div_class.mod_less)
+apply(transfer, fact semiring_numeral_div_class.div_positive)
+apply(transfer, fact semiring_numeral_div_class.mod_less_eq_dividend)
+apply(transfer, fact semiring_numeral_div_class.pos_mod_bound)
+apply(transfer, fact semiring_numeral_div_class.pos_mod_sign)
+apply(transfer, fact semiring_numeral_div_class.mod_mult2_eq)
+apply(transfer, fact semiring_numeral_div_class.div_mult2_eq)
+apply(transfer, fact discrete)
done
subsection {* Finite class *}
--- a/src/HOL/NSA/transfer.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/NSA/transfer.ML Mon Mar 30 11:36:30 2015 +0200
@@ -68,7 +68,7 @@
in Goal.prove ctxt [] [] (Logic.mk_equals (t,u)) (K tac) end
fun transfer_tac ctxt ths =
- SUBGOAL (fn (t,i) =>
+ SUBGOAL (fn (t, _) =>
(fn th =>
let
val tr = transfer_thm_of ctxt ths t
--- a/src/HOL/NanoJava/Equivalence.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/NanoJava/Equivalence.thy Mon Mar 30 11:36:30 2015 +0200
@@ -105,8 +105,8 @@
apply (rule hoare_ehoare.induct)
(*18*)
apply (tactic {* ALLGOALS (REPEAT o dresolve_tac @{context} [@{thm all_conjunct2}, @{thm all3_conjunct2}]) *})
-apply (tactic {* ALLGOALS (REPEAT o Rule_Insts.thin_tac @{context} "hoare ?x ?y" []) *})
-apply (tactic {* ALLGOALS (REPEAT o Rule_Insts.thin_tac @{context} "ehoare ?x ?y" []) *})
+apply (tactic {* ALLGOALS (REPEAT o Rule_Insts.thin_tac @{context} "hoare _ _" []) *})
+apply (tactic {* ALLGOALS (REPEAT o Rule_Insts.thin_tac @{context} "ehoare _ _" []) *})
apply (simp_all only: cnvalid1_eq cenvalid_def2)
apply fast
apply fast
--- a/src/HOL/Nat.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Nat.thy Mon Mar 30 11:36:30 2015 +0200
@@ -246,11 +246,10 @@
fix n m q :: nat
show "(n + m) + q = n + (m + q)" by (induct n) simp_all
show "n + m = m + n" by (induct n) simp_all
+ show "m + n - m = n" by (induct m) simp_all
+ show "n - m - q = n - (m + q)" by (induct q) (simp_all add: diff_Suc)
show "0 + n = n" by simp
- show "n - 0 = n" by simp
show "0 - n = 0" by simp
- show "(q + n) - (q + m) = n - m" by (induct q) simp_all
- show "n - m - q = n - (m + q)" by (induct q) (simp_all add: diff_Suc)
qed
end
@@ -283,7 +282,6 @@
show "n * m = m * n" by (induct n) simp_all
show "(n * m) * q = n * (m * q)" by (induct n) (simp_all add: add_mult_distrib)
show "(n + m) * q = n * q + m * q" by (rule add_mult_distrib)
- assume "n + m = n + q" thus "m = q" by (induct n) simp_all
qed
end
@@ -355,7 +353,7 @@
by (fact add_diff_cancel_right')
lemma diff_cancel: "(k + m) - (k + n) = m - (n::nat)"
- by (fact comm_monoid_diff_class.add_diff_cancel_left)
+ by (fact add_diff_cancel_left)
lemma diff_cancel2: "(m + k) - (n + k) = m - (n::nat)"
by (fact add_diff_cancel_right)
@@ -368,12 +366,20 @@
text {* Difference distributes over multiplication *}
-lemma diff_mult_distrib: "((m::nat) - n) * k = (m * k) - (n * k)"
-by (induct m n rule: diff_induct) (simp_all add: diff_cancel)
-
-lemma diff_mult_distrib2: "k * ((m::nat) - n) = (k * m) - (k * n)"
-by (simp add: diff_mult_distrib mult.commute [of k])
- -- {* NOT added as rewrites, since sometimes they are used from right-to-left *}
+instance nat :: comm_semiring_1_diff_distrib
+proof
+ fix k m n :: nat
+ show "k * ((m::nat) - n) = (k * m) - (k * n)"
+ by (induct m n rule: diff_induct) simp_all
+qed
+
+lemma diff_mult_distrib:
+ "((m::nat) - n) * k = (m * k) - (n * k)"
+ by (fact left_diff_distrib')
+
+lemma diff_mult_distrib2:
+ "k * ((m::nat) - n) = (k * m) - (k * n)"
+ by (fact right_diff_distrib')
subsubsection {* Multiplication *}
@@ -777,18 +783,13 @@
apply (simp_all add: add_less_mono)
done
-text{*The naturals form an ordered @{text comm_semiring_1_cancel}*}
+text{*The naturals form an ordered @{text semidom}*}
instance nat :: linordered_semidom
proof
show "0 < (1::nat)" by simp
show "\<And>m n q :: nat. m \<le> n \<Longrightarrow> q + m \<le> q + n" by simp
show "\<And>m n q :: nat. m < n \<Longrightarrow> 0 < q \<Longrightarrow> q * m < q * n" by (simp add: mult_less_mono2)
-qed
-
-instance nat :: semiring_no_zero_divisors
-proof
- fix m n :: nat
- show "m \<noteq> 0 \<Longrightarrow> n \<noteq> 0 \<Longrightarrow> m * n \<noteq> 0" by simp
+ show "\<And>m n :: nat. m \<noteq> 0 \<Longrightarrow> n \<noteq> 0 \<Longrightarrow> m * n \<noteq> 0" by simp
qed
@@ -1878,48 +1879,6 @@
subsection {* The divides relation on @{typ nat} *}
-instance nat :: semiring_dvd
-proof
- fix m n q :: nat
- show "m dvd q * m + n \<longleftrightarrow> m dvd n" (is "?P \<longleftrightarrow> ?Q")
- proof
- assume ?Q then show ?P by simp
- next
- assume ?P then obtain d where *: "q * m + n = m * d" ..
- show ?Q
- proof (cases "n = 0")
- case True then show ?Q by simp
- next
- case False
- with * have "q * m < m * d"
- using less_add_eq_less [of 0 n "q * m" "m * d"] by simp
- then have "q \<le> d" by (auto intro: ccontr simp add: mult.commute [of m])
- with * [symmetric] have "n = m * (d - q)"
- by (simp add: diff_mult_distrib2 mult.commute [of m])
- then show ?Q ..
- qed
- qed
- assume "m dvd n + q" and "m dvd n"
- show "m dvd q"
- proof -
- from `m dvd n` obtain d where "n = m * d" ..
- moreover from `m dvd n + q` obtain e where "n + q = m * e" ..
- ultimately have *: "m * d + q = m * e" by simp
- show "m dvd q"
- proof (cases "q = 0")
- case True then show ?thesis by simp
- next
- case False
- with * have "m * d < m * e"
- using less_add_eq_less [of 0 q "m * d" "m * e"] by simp
- then have "d \<le> e" by (auto intro: ccontr)
- with * have "q = m * (e - d)"
- by (simp add: diff_mult_distrib2)
- then show ?thesis ..
- qed
- qed
-qed
-
lemma dvd_1_left [iff]: "Suc 0 dvd k"
unfolding dvd_def by simp
--- a/src/HOL/Nominal/Examples/Standardization.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Nominal/Examples/Standardization.thy Mon Mar 30 11:36:30 2015 +0200
@@ -384,7 +384,7 @@
apply (subgoal_tac "r = (Lam [xa]. [(xa, x)] \<bullet> s)")
prefer 2
apply (simp add: calc_atm)
- apply (thin_tac "r = ?t")
+ apply (thin_tac "r = _")
apply simp
apply (rule exI)
apply (rule conjI)
--- a/src/HOL/Nominal/nominal_atoms.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Nominal/nominal_atoms.ML Mon Mar 30 11:36:30 2015 +0200
@@ -131,7 +131,7 @@
(HOLogic.mk_exists ("x",@{typ nat},HOLogic.mk_eq (y,Const (ak_sign,inj_type) $ Bound 0)))
val proof2 = fn {prems, context = ctxt} =>
- Induct_Tacs.case_tac ctxt "y" [] 1 THEN
+ Induct_Tacs.case_tac ctxt "y" [] NONE 1 THEN
asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp1) 1 THEN
rtac @{thm exI} 1 THEN
rtac @{thm refl} 1
--- a/src/HOL/Number_Theory/Cong.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Number_Theory/Cong.thy Mon Mar 30 11:36:30 2015 +0200
@@ -443,7 +443,7 @@
apply (erule ssubst)
apply (subst zmod_zmult2_eq)
apply (auto simp add: mod_add_left_eq mod_minus_right div_minus_right)
- apply (metis mod_diff_left_eq mod_diff_right_eq mod_mult_self1_is_0 semiring_numeral_div_class.diff_zero)+
+ apply (metis mod_diff_left_eq mod_diff_right_eq mod_mult_self1_is_0 diff_zero)+
done
lemma cong_to_1_nat: "([(a::nat) = 1] (mod n)) \<Longrightarrow> (n dvd (a - 1))"
--- a/src/HOL/Parity.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Parity.thy Mon Mar 30 11:36:30 2015 +0200
@@ -11,13 +11,15 @@
subsection {* Ring structures with parity and @{text even}/@{text odd} predicates *}
-class semiring_parity = semiring_dvd + semiring_numeral +
+class semiring_parity = comm_semiring_1_diff_distrib + numeral +
assumes odd_one [simp]: "\<not> 2 dvd 1"
assumes odd_even_add: "\<not> 2 dvd a \<Longrightarrow> \<not> 2 dvd b \<Longrightarrow> 2 dvd a + b"
assumes even_multD: "2 dvd a * b \<Longrightarrow> 2 dvd a \<or> 2 dvd b"
assumes odd_ex_decrement: "\<not> 2 dvd a \<Longrightarrow> \<exists>b. a = b + 1"
begin
+subclass semiring_numeral ..
+
abbreviation even :: "'a \<Rightarrow> bool"
where
"even a \<equiv> 2 dvd a"
@@ -97,9 +99,11 @@
end
-class ring_parity = comm_ring_1 + semiring_parity
+class ring_parity = ring + semiring_parity
begin
+subclass comm_ring_1 ..
+
lemma even_minus [simp]:
"even (- a) \<longleftrightarrow> even a"
by (fact dvd_minus_iff)
--- a/src/HOL/Probability/measurable.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Probability/measurable.ML Mon Mar 30 11:36:30 2015 +0200
@@ -147,8 +147,9 @@
| indep (Abs (_,_,t)) top bot = indep t (top + 1) (bot + 1)
| indep _ _ _ = true;
-fun cnt_prefixes ctxt (Abs (n, T, t)) = let
- fun is_countable t = Type.of_sort (Proof_Context.tsig_of ctxt) (t, @{sort countable})
+fun cnt_prefixes ctxt (Abs (n, T, t)) =
+ let
+ fun is_countable ty = Sign.of_sort (Proof_Context.theory_of ctxt) (ty, @{sort countable})
fun cnt_walk (Abs (ns, T, t)) Ts =
map (fn (t', t'') => (Abs (ns, T, t'), t'')) (cnt_walk t (T::Ts))
| cnt_walk (f $ g) Ts = let
--- a/src/HOL/ROOT Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/ROOT Mon Mar 30 11:36:30 2015 +0200
@@ -824,7 +824,7 @@
theories SPARK
session "HOL-SPARK-Examples" in "SPARK/Examples" = "HOL-SPARK" +
- options [document = false]
+ options [document = false, spark_prv = false]
theories
"Gcd/Greatest_Common_Divisor"
@@ -877,7 +877,7 @@
"RIPEMD-160/rmd/s_r.siv"
session "HOL-SPARK-Manual" in "SPARK/Manual" = "HOL-SPARK" +
- options [show_question_marks = false]
+ options [show_question_marks = false, spark_prv = false]
theories
Example_Verification
VC_Principles
--- a/src/HOL/Rings.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Rings.thy Mon Mar 30 11:36:30 2015 +0200
@@ -197,7 +197,6 @@
end
class comm_semiring_1 = zero_neq_one + comm_semiring_0 + comm_monoid_mult
- (*previously almost_semiring*)
begin
subclass semiring_1 ..
@@ -228,51 +227,6 @@
end
-class semiring_dvd = comm_semiring_1 +
- assumes dvd_add_times_triv_left_iff [simp]: "a dvd c * a + b \<longleftrightarrow> a dvd b"
- assumes dvd_addD: "a dvd b + c \<Longrightarrow> a dvd b \<Longrightarrow> a dvd c"
-begin
-
-lemma dvd_add_times_triv_right_iff [simp]:
- "a dvd b + c * a \<longleftrightarrow> a dvd b"
- using dvd_add_times_triv_left_iff [of a c b] by (simp add: ac_simps)
-
-lemma dvd_add_triv_left_iff [simp]:
- "a dvd a + b \<longleftrightarrow> a dvd b"
- using dvd_add_times_triv_left_iff [of a 1 b] by simp
-
-lemma dvd_add_triv_right_iff [simp]:
- "a dvd b + a \<longleftrightarrow> a dvd b"
- using dvd_add_times_triv_right_iff [of a b 1] by simp
-
-lemma dvd_add_right_iff:
- assumes "a dvd b"
- shows "a dvd b + c \<longleftrightarrow> a dvd c"
- using assms by (auto dest: dvd_addD)
-
-lemma dvd_add_left_iff:
- assumes "a dvd c"
- shows "a dvd b + c \<longleftrightarrow> a dvd b"
- using assms dvd_add_right_iff [of a c b] by (simp add: ac_simps)
-
-end
-
-class no_zero_divisors = zero + times +
- assumes no_zero_divisors: "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> a * b \<noteq> 0"
-begin
-
-lemma divisors_zero:
- assumes "a * b = 0"
- shows "a = 0 \<or> b = 0"
-proof (rule classical)
- assume "\<not> (a = 0 \<or> b = 0)"
- then have "a \<noteq> 0" and "b \<noteq> 0" by auto
- with no_zero_divisors have "a * b \<noteq> 0" by blast
- with assms show ?thesis by simp
-qed
-
-end
-
class semiring_1_cancel = semiring + cancel_comm_monoid_add
+ zero_neq_one + monoid_mult
begin
@@ -293,6 +247,65 @@
end
+class comm_semiring_1_diff_distrib = comm_semiring_1_cancel +
+ assumes right_diff_distrib' [algebra_simps]: "a * (b - c) = a * b - a * c"
+begin
+
+lemma left_diff_distrib' [algebra_simps]:
+ "(b - c) * a = b * a - c * a"
+ by (simp add: algebra_simps)
+
+lemma dvd_add_times_triv_left_iff [simp]:
+ "a dvd c * a + b \<longleftrightarrow> a dvd b"
+proof -
+ have "a dvd a * c + b \<longleftrightarrow> a dvd b" (is "?P \<longleftrightarrow> ?Q")
+ proof
+ assume ?Q then show ?P by simp
+ next
+ assume ?P
+ then obtain d where "a * c + b = a * d" ..
+ then have "a * c + b - a * c = a * d - a * c" by simp
+ then have "b = a * d - a * c" by simp
+ then have "b = a * (d - c)" by (simp add: algebra_simps)
+ then show ?Q ..
+ qed
+ then show "a dvd c * a + b \<longleftrightarrow> a dvd b" by (simp add: ac_simps)
+qed
+
+lemma dvd_add_times_triv_right_iff [simp]:
+ "a dvd b + c * a \<longleftrightarrow> a dvd b"
+ using dvd_add_times_triv_left_iff [of a c b] by (simp add: ac_simps)
+
+lemma dvd_add_triv_left_iff [simp]:
+ "a dvd a + b \<longleftrightarrow> a dvd b"
+ using dvd_add_times_triv_left_iff [of a 1 b] by simp
+
+lemma dvd_add_triv_right_iff [simp]:
+ "a dvd b + a \<longleftrightarrow> a dvd b"
+ using dvd_add_times_triv_right_iff [of a b 1] by simp
+
+lemma dvd_add_right_iff:
+ assumes "a dvd b"
+ shows "a dvd b + c \<longleftrightarrow> a dvd c" (is "?P \<longleftrightarrow> ?Q")
+proof
+ assume ?P then obtain d where "b + c = a * d" ..
+ moreover from `a dvd b` obtain e where "b = a * e" ..
+ ultimately have "a * e + c = a * d" by simp
+ then have "a * e + c - a * e = a * d - a * e" by simp
+ then have "c = a * d - a * e" by simp
+ then have "c = a * (d - e)" by (simp add: algebra_simps)
+ then show ?Q ..
+next
+ assume ?Q with assms show ?P by simp
+qed
+
+lemma dvd_add_left_iff:
+ assumes "a dvd c"
+ shows "a dvd b + c \<longleftrightarrow> a dvd b"
+ using assms dvd_add_right_iff [of a c b] by (simp add: ac_simps)
+
+end
+
class ring = semiring + ab_group_add
begin
@@ -364,33 +377,13 @@
end
class comm_ring_1 = comm_ring + zero_neq_one + comm_monoid_mult
- (*previously ring*)
begin
subclass ring_1 ..
subclass comm_semiring_1_cancel ..
-subclass semiring_dvd
-proof
- fix a b c
- show "a dvd c * a + b \<longleftrightarrow> a dvd b" (is "?P \<longleftrightarrow> ?Q")
- proof
- assume ?Q then show ?P by simp
- next
- assume ?P then obtain d where "c * a + b = a * d" ..
- then have "b = a * (d - c)" by (simp add: algebra_simps)
- then show ?Q ..
- qed
- assume "a dvd b + c" and "a dvd b"
- show "a dvd c"
- proof -
- from `a dvd b` obtain d where "b = a * d" ..
- moreover from `a dvd b + c` obtain e where "b + c = a * e" ..
- ultimately have "a * d + c = a * e" by simp
- then have "c = a * (e - d)" by (simp add: algebra_simps)
- then show "a dvd c" ..
- qed
-qed
+subclass comm_semiring_1_diff_distrib
+ by unfold_locales (simp add: algebra_simps)
lemma dvd_minus_iff [simp]: "x dvd - y \<longleftrightarrow> x dvd y"
proof
@@ -422,9 +415,20 @@
end
-class semiring_no_zero_divisors = semiring_0 + no_zero_divisors
+class semiring_no_zero_divisors = semiring_0 +
+ assumes no_zero_divisors: "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> a * b \<noteq> 0"
begin
+lemma divisors_zero:
+ assumes "a * b = 0"
+ shows "a = 0 \<or> b = 0"
+proof (rule classical)
+ assume "\<not> (a = 0 \<or> b = 0)"
+ then have "a \<noteq> 0" and "b \<noteq> 0" by auto
+ with no_zero_divisors have "a * b \<noteq> 0" by blast
+ with assms show ?thesis by simp
+qed
+
lemma mult_eq_0_iff [simp]:
shows "a * b = 0 \<longleftrightarrow> a = 0 \<or> b = 0"
proof (cases "a = 0 \<or> b = 0")
@@ -498,23 +502,15 @@
end
-class idom = comm_ring_1 + no_zero_divisors
+class semidom = comm_semiring_1_cancel + semiring_no_zero_divisors
+
+class idom = comm_ring_1 + semiring_no_zero_divisors
begin
+subclass semidom ..
+
subclass ring_1_no_zero_divisors ..
-lemma square_eq_iff: "a * a = b * b \<longleftrightarrow> (a = b \<or> a = - b)"
-proof
- assume "a * a = b * b"
- then have "(a - b) * (a + b) = 0"
- by (simp add: algebra_simps)
- then show "a = b \<or> a = - b"
- by (simp add: eq_neg_iff_add_eq_0)
-next
- assume "a = b \<or> a = - b"
- then show "a * a = b * b" by auto
-qed
-
lemma dvd_mult_cancel_right [simp]:
"a * c dvd b * c \<longleftrightarrow> c = 0 \<or> a dvd b"
proof -
@@ -535,6 +531,18 @@
finally show ?thesis .
qed
+lemma square_eq_iff: "a * a = b * b \<longleftrightarrow> (a = b \<or> a = - b)"
+proof
+ assume "a * a = b * b"
+ then have "(a - b) * (a + b) = 0"
+ by (simp add: algebra_simps)
+ then show "a = b \<or> a = - b"
+ by (simp add: eq_neg_iff_add_eq_0)
+next
+ assume "a = b \<or> a = - b"
+ then show "a * a = b * b" by auto
+qed
+
end
text {*
@@ -846,9 +854,6 @@
end
-(* The "strict" suffix can be seen as describing the combination of linordered_ring and no_zero_divisors.
- Basically, linordered_ring + no_zero_divisors = linordered_ring_strict.
- *)
class linordered_ring_strict = ring + linordered_semiring_strict
+ ordered_ab_group_add + abs_if
begin
@@ -994,8 +999,7 @@
end
-class linordered_semidom = comm_semiring_1_cancel + linordered_comm_semiring_strict +
- (*previously linordered_semiring*)
+class linordered_semidom = semidom + linordered_comm_semiring_strict +
assumes zero_less_one [simp]: "0 < 1"
begin
@@ -1029,7 +1033,6 @@
class linordered_idom = comm_ring_1 +
linordered_comm_semiring_strict + ordered_ab_group_add +
abs_if + sgn_if
- (*previously linordered_ring*)
begin
subclass linordered_semiring_1_strict ..
@@ -1195,7 +1198,7 @@
lemma mult_right_le_one_le:
"0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> y \<le> 1 \<Longrightarrow> x * y \<le> x"
- by (auto simp add: mult_le_cancel_left2)
+ by (rule mult_left_le)
lemma mult_left_le_one_le:
"0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> y \<le> 1 \<Longrightarrow> y * x \<le> x"
@@ -1265,4 +1268,3 @@
code_module Rings \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
end
-
--- a/src/HOL/SET_Protocol/Cardholder_Registration.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/SET_Protocol/Cardholder_Registration.thy Mon Mar 30 11:36:30 2015 +0200
@@ -771,7 +771,7 @@
==> Nonce N \<notin> analz (knows Spy evs) -->
(\<exists>k i W. Says (Cardholder k) (CA i) {|U,W|} \<in> set evs &
Cardholder k \<notin> bad & CA i \<notin> bad)"
-apply (erule_tac P = "U \<in> ?H" in rev_mp)
+apply (erule_tac P = "U \<in> H" for H in rev_mp)
apply (erule set_cr.induct)
apply (valid_certificate_tac [8]) --{*for message 5*}
apply (simp_all del: image_insert image_Un imp_disjL
--- a/src/HOL/SPARK/Tools/spark_vcs.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/SPARK/Tools/spark_vcs.ML Mon Mar 30 11:36:30 2015 +0200
@@ -953,29 +953,33 @@
| x => x);
fun close incomplete thy =
- thy |>
- VCs.map (fn
- {pfuns, type_map, env = SOME {vcs, path, ...}} =>
+ thy |> VCs.map (fn {pfuns, type_map, env} =>
+ (case env of
+ NONE => error "No SPARK environment is currently open"
+ | SOME {vcs, path, ...} =>
let
val (proved, unproved) = partition_vcs vcs;
val _ = Thm.join_proofs (maps (#2 o snd) proved);
- val (proved', proved'') = List.partition (fn (_, (_, thms, _, _)) =>
- exists (#oracle o Thm.peek_status) thms) proved
- in
- (if null unproved then ()
- else (if incomplete then warning else error)
- (Pretty.string_of (pp_open_vcs unproved));
- if null proved' then ()
- else warning (Pretty.string_of (pp_vcs
+ val (proved', proved'') =
+ List.partition (fn (_, (_, thms, _, _)) =>
+ exists (#oracle o Thm.peek_status) thms) proved;
+ val _ =
+ if null unproved then ()
+ else (if incomplete then warning else error) (Pretty.string_of (pp_open_vcs unproved));
+ val _ =
+ if null proved' then ()
+ else warning (Pretty.string_of (pp_vcs
"The following VCs are not marked as proved \
\because their proofs contain oracles:" proved'));
- File.write (Path.ext "prv" path)
- (implode (map (fn (s, _) => snd (strip_number s) ^
- " -- proved by " ^ Distribution.version ^ "\n") proved''));
- {pfuns = pfuns, type_map = type_map, env = NONE})
- end
- | _ => error "No SPARK environment is currently open") |>
- Sign.parent_path;
+ val prv_path = Path.ext "prv" path;
+ val _ =
+ if File.exists prv_path orelse Options.default_bool @{system_option spark_prv} then
+ File.write prv_path
+ (implode (map (fn (s, _) => snd (strip_number s) ^
+ " -- proved by " ^ Distribution.version ^ "\n") proved''))
+ else ();
+ in {pfuns = pfuns, type_map = type_map, env = NONE} end))
+ |> Sign.parent_path;
(** set up verification conditions **)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/etc/options Mon Mar 30 11:36:30 2015 +0200
@@ -0,0 +1,6 @@
+(* :mode=isabelle-options: *)
+
+section "HOL-SPARK"
+
+option spark_prv : bool = true
+ -- "produce proof review file after 'spark_end'"
--- a/src/HOL/TLA/Memory/MemoryImplementation.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/TLA/Memory/MemoryImplementation.thy Mon Mar 30 11:36:30 2015 +0200
@@ -798,7 +798,7 @@
fun split_idle_tac ctxt =
SELECT_GOAL
(TRY (rtac @{thm actionI} 1) THEN
- Induct_Tacs.case_tac ctxt "(s,t) |= unchanged (e p, c p, r p, m p, rmhist!p)" [] 1 THEN
+ Induct_Tacs.case_tac ctxt "(s,t) |= unchanged (e p, c p, r p, m p, rmhist!p)" [] NONE 1 THEN
rewrite_goals_tac ctxt @{thms action_rews} THEN
forward_tac ctxt [temp_use ctxt @{thm Step1_4_7}] 1 THEN
asm_full_simp_tac ctxt 1);
--- a/src/HOL/Tools/BNF/bnf_comp.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Tools/BNF/bnf_comp.ML Mon Mar 30 11:36:30 2015 +0200
@@ -791,7 +791,10 @@
val repTA = mk_T_of_bnf Ds As bnf;
val T_bind = qualify b;
- val TA_params = Term.add_tfreesT repTA (if has_live_phantoms then As' else []);
+ val all_TA_params_in_order = Term.add_tfreesT repTA As';
+ val TA_params =
+ (if has_live_phantoms then all_TA_params_in_order
+ else inter (op =) (Term.add_tfreesT repTA []) all_TA_params_in_order);
val ((TA, (Rep_name, Abs_name, type_definition, Abs_inverse, Abs_inject, _)), lthy) =
maybe_typedef lthy has_live_phantoms (T_bind, TA_params, NoSyn) (HOLogic.mk_UNIV repTA) NONE
(fn _ => EVERY' [rtac exI, rtac UNIV_I] 1) lthy;
@@ -906,7 +909,9 @@
val notes = (map_b, map_def) :: (rel_b, rel_def) :: (set_bs ~~ set_defs)
|> map (fn (b, def) => ((b, []), [([def], [])]))
- val (noted, lthy'') = lthy' |> Local_Theory.notes notes
+ val (noted, lthy'') = lthy'
+ |> Local_Theory.notes notes
+ ||> (if repTA = TA then I else register_bnf_raw (fst (dest_Type TA)) bnf'')
in
((morph_bnf (substitute_noted_thm noted) bnf'', (all_deads, absT_info)), lthy'')
end;
--- a/src/HOL/Tools/BNF/bnf_def.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Tools/BNF/bnf_def.ML Mon Mar 30 11:36:30 2015 +0200
@@ -1609,7 +1609,8 @@
[Pretty.block [Pretty.str (bdN ^ ":"), Pretty.brk 1,
Pretty.quote (Syntax.pretty_term ctxt bd)]]);
in
- Pretty.big_list "BNFs:" (map pretty_bnf (Symtab.dest (Data.get (Context.Proof ctxt))))
+ Pretty.big_list "Registered bounded natural functors:"
+ (map pretty_bnf (sort_wrt fst (Symtab.dest (Data.get (Context.Proof ctxt)))))
|> Pretty.writeln
end;
--- a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML Mon Mar 30 11:36:30 2015 +0200
@@ -83,6 +83,8 @@
val flat_corec_preds_predsss_gettersss: 'a list -> 'a list list list -> 'a list list list ->
'a list
+ val mk_ctor: typ list -> term -> term
+ val mk_dtor: typ list -> term -> term
val liveness_of_fp_bnf: int -> BNF_Def.bnf -> bool list
val nesting_bnfs: Proof.context -> typ list list list -> typ list -> BNF_Def.bnf list
@@ -551,7 +553,7 @@
b_names Ts thmss)
#> filter_out (null o fst o hd o snd);
-fun derive_map_set_rel_thms plugins fp n live As Bs abs_inverses ctr_defs' fp_nesting_set_maps
+fun derive_map_set_rel_thms plugins fp live As Bs abs_inverses ctr_defs' fp_nesting_set_maps
live_nesting_map_id0s live_nesting_set_maps live_nesting_rel_eqs fp_b_name fp_bnf fpT ctor
ctor_dtor dtor_ctor pre_map_def pre_set_defs pre_rel_def fp_map_thm fp_set_thms fp_rel_thm
ctr_Tss abs
@@ -2091,11 +2093,11 @@
val ((xtor_co_recs, recs_args_types, corecs_args_types), lthy') =
mk_co_recs_prelims fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_recs0 lthy;
- fun define_ctrs_dtrs_for_type fp (((((((((((((((((((((((((((fp_bnf, fp_b), fpT), ctor), dtor),
+ fun define_ctrs_dtrs_for_type fp ((((((((((((((((((((((((((fp_bnf, fp_b), fpT), ctor), dtor),
xtor_co_rec), ctor_dtor), dtor_ctor), ctor_inject), pre_map_def), pre_set_defs),
pre_rel_def), fp_map_thm), fp_set_thms), fp_rel_thm), n), ks), ms), abs),
- abs_inject), abs_inverse), type_definition), ctr_bindings), ctr_mixfixes), ctr_Tss),
- disc_bindings), sel_bindingss), raw_sel_default_eqs) no_defs_lthy =
+ abs_inject), type_definition), ctr_bindings), ctr_mixfixes), ctr_Tss), disc_bindings),
+ sel_bindingss), raw_sel_default_eqs) no_defs_lthy =
let
val fp_b_name = Binding.name_of fp_b;
@@ -2195,7 +2197,7 @@
in
(wrap_ctrs
#> (fn (ctr_sugar, lthy) =>
- derive_map_set_rel_thms plugins fp n live As Bs abs_inverses ctr_defs'
+ derive_map_set_rel_thms plugins fp live As Bs abs_inverses ctr_defs'
fp_nesting_set_maps live_nesting_map_id0s live_nesting_set_maps live_nesting_rel_eqs
fp_b_name fp_bnf fpT ctor ctor_dtor dtor_ctor pre_map_def pre_set_defs pre_rel_def
fp_map_thm fp_set_thms fp_rel_thm ctr_Tss abs ctr_sugar lthy
@@ -2255,8 +2257,7 @@
val gmaps = map (fn map0 => Term.list_comb (mk_map live As Bs map0, live_gs)) maps0;
- val rec_Ts as rec_T1 :: _ = map fastype_of recs;
- val rec_arg_Ts = binder_fun_types rec_T1;
+ val rec_arg_Ts = binder_fun_types (hd (map fastype_of recs));
val B_ify_T = Term.typ_subst_atomic (As ~~ Bs);
@@ -2407,8 +2408,7 @@
val fmaps = map (fn map0 => Term.list_comb (mk_map live As Bs map0, live_fs)) maps0;
- val corec_Ts as corec_T1 :: _ = map fastype_of corecs;
- val corec_arg_Ts = binder_fun_types corec_T1;
+ val corec_arg_Ts = binder_fun_types (hd (map fastype_of corecs));
val B_ify = Term.subst_atomic_types (As ~~ Bs);
val B_ify_T = Term.typ_subst_atomic (As ~~ Bs);
@@ -2564,8 +2564,8 @@
|> fold_map (define_ctrs_dtrs_for_type fp) (fp_bnfs ~~ fp_bs ~~ fpTs ~~ ctors ~~ dtors ~~
xtor_co_recs ~~ ctor_dtors ~~ dtor_ctors ~~ ctor_injects ~~ pre_map_defs ~~ pre_set_defss ~~
pre_rel_defs ~~ xtor_maps ~~ xtor_setss ~~ xtor_rels ~~ ns ~~ kss ~~ mss ~~
- abss ~~ abs_injects ~~ abs_inverses ~~ type_definitions ~~ ctr_bindingss ~~ ctr_mixfixess ~~
- ctr_Tsss ~~ disc_bindingss ~~ sel_bindingsss ~~ raw_sel_default_eqss)
+ abss ~~ abs_injects ~~ type_definitions ~~ ctr_bindingss ~~ ctr_mixfixess ~~ ctr_Tsss ~~
+ disc_bindingss ~~ sel_bindingsss ~~ raw_sel_default_eqss)
|> wrap_ctrs_derive_map_set_rel_thms_define_co_rec_for_types
|> case_fp fp derive_note_induct_recs_thms_for_types
derive_note_coinduct_corecs_thms_for_types;
--- a/src/HOL/Tools/BNF/bnf_fp_n2m.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Tools/BNF/bnf_fp_n2m.ML Mon Mar 30 11:36:30 2015 +0200
@@ -478,7 +478,8 @@
used by "primrec", "primcorecursive", and "datatype_compat". *)
val fp_res =
({Ts = fpTs, bnfs = of_fp_res #bnfs, dtors = dtors, ctors = ctors,
- xtor_co_recs = co_recs, xtor_co_induct = xtor_co_induct_thm,
+ xtor_un_folds = co_recs (*wrong*), xtor_co_recs = co_recs,
+ xtor_co_induct = xtor_co_induct_thm,
dtor_ctors = of_fp_res #dtor_ctors (*too general types*),
ctor_dtors = of_fp_res #ctor_dtors (*too general types*),
ctor_injects = of_fp_res #ctor_injects (*too general types*),
@@ -486,11 +487,9 @@
xtor_maps = of_fp_res #xtor_maps (*too general types and terms*),
xtor_setss = of_fp_res #xtor_setss (*too general types and terms*),
xtor_rels = of_fp_res #xtor_rels (*too general types and terms*),
- xtor_co_rec_thms = xtor_co_rec_thms,
+ xtor_un_fold_thms = xtor_co_rec_thms (*wrong*), xtor_co_rec_thms = xtor_co_rec_thms,
xtor_co_rec_o_maps = fp_rec_o_maps (*theorems about old constants*),
- xtor_rel_co_induct = xtor_rel_co_induct,
- dtor_set_inducts = [],
- xtor_co_rec_transfers = []}
+ xtor_rel_co_induct = xtor_rel_co_induct, dtor_set_inducts = [], xtor_co_rec_transfers = []}
|> morph_fp_result (Morphism.term_morphism "BNF" (singleton (Variable.polymorphic lthy))));
in
(fp_res, lthy)
--- a/src/HOL/Tools/BNF/bnf_fp_util.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Tools/BNF/bnf_fp_util.ML Mon Mar 30 11:36:30 2015 +0200
@@ -15,6 +15,7 @@
bnfs: BNF_Def.bnf list,
ctors: term list,
dtors: term list,
+ xtor_un_folds: term list,
xtor_co_recs: term list,
xtor_co_induct: thm,
dtor_ctors: thm list,
@@ -24,6 +25,7 @@
xtor_maps: thm list,
xtor_setss: thm list list,
xtor_rels: thm list,
+ xtor_un_fold_thms: thm list,
xtor_co_rec_thms: thm list,
xtor_co_rec_o_maps: thm list,
xtor_rel_co_induct: thm,
@@ -208,6 +210,7 @@
bnfs: BNF_Def.bnf list,
ctors: term list,
dtors: term list,
+ xtor_un_folds: term list,
xtor_co_recs: term list,
xtor_co_induct: thm,
dtor_ctors: thm list,
@@ -217,20 +220,22 @@
xtor_maps: thm list,
xtor_setss: thm list list,
xtor_rels: thm list,
+ xtor_un_fold_thms: thm list,
xtor_co_rec_thms: thm list,
xtor_co_rec_o_maps: thm list,
xtor_rel_co_induct: thm,
dtor_set_inducts: thm list,
xtor_co_rec_transfers: thm list};
-fun morph_fp_result phi {Ts, bnfs, ctors, dtors, xtor_co_recs, xtor_co_induct,
- dtor_ctors, ctor_dtors, ctor_injects, dtor_injects, xtor_maps, xtor_setss,
- xtor_rels, xtor_co_rec_thms, xtor_co_rec_o_maps, xtor_rel_co_induct,
- dtor_set_inducts, xtor_co_rec_transfers} =
+fun morph_fp_result phi {Ts, bnfs, ctors, dtors, xtor_un_folds, xtor_co_recs, xtor_co_induct,
+ dtor_ctors, ctor_dtors, ctor_injects, dtor_injects, xtor_maps, xtor_setss, xtor_rels,
+ xtor_un_fold_thms, xtor_co_rec_thms, xtor_co_rec_o_maps, xtor_rel_co_induct, dtor_set_inducts,
+ xtor_co_rec_transfers} =
{Ts = map (Morphism.typ phi) Ts,
bnfs = map (morph_bnf phi) bnfs,
ctors = map (Morphism.term phi) ctors,
dtors = map (Morphism.term phi) dtors,
+ xtor_un_folds = map (Morphism.term phi) xtor_un_folds,
xtor_co_recs = map (Morphism.term phi) xtor_co_recs,
xtor_co_induct = Morphism.thm phi xtor_co_induct,
dtor_ctors = map (Morphism.thm phi) dtor_ctors,
@@ -240,6 +245,7 @@
xtor_maps = map (Morphism.thm phi) xtor_maps,
xtor_setss = map (map (Morphism.thm phi)) xtor_setss,
xtor_rels = map (Morphism.thm phi) xtor_rels,
+ xtor_un_fold_thms = map (Morphism.thm phi) xtor_un_fold_thms,
xtor_co_rec_thms = map (Morphism.thm phi) xtor_co_rec_thms,
xtor_co_rec_o_maps = map (Morphism.thm phi) xtor_co_rec_o_maps,
xtor_rel_co_induct = Morphism.thm phi xtor_rel_co_induct,
--- a/src/HOL/Tools/BNF/bnf_gfp.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Tools/BNF/bnf_gfp.ML Mon Mar 30 11:36:30 2015 +0200
@@ -2538,13 +2538,13 @@
val lthy' = lthy |> note_all ? snd o Local_Theory.notes (common_notes @ notes @ Jbnf_notes);
val fp_res =
- {Ts = Ts, bnfs = Jbnfs, ctors = ctors, dtors = dtors, xtor_co_recs = corecs,
- xtor_co_induct = dtor_coinduct_thm, dtor_ctors = dtor_ctor_thms, ctor_dtors = ctor_dtor_thms,
- ctor_injects = ctor_inject_thms, dtor_injects = dtor_inject_thms,
- xtor_maps = dtor_Jmap_thms, xtor_setss = dtor_Jset_thmss',
- xtor_rels = dtor_Jrel_thms, xtor_co_rec_thms = dtor_corec_thms,
- xtor_co_rec_o_maps = dtor_corec_o_Jmap_thms, xtor_rel_co_induct = Jrel_coinduct_thm,
- dtor_set_inducts = dtor_Jset_induct_thms,
+ {Ts = Ts, bnfs = Jbnfs, ctors = ctors, dtors = dtors, xtor_un_folds = unfolds,
+ xtor_co_recs = corecs, xtor_co_induct = dtor_coinduct_thm, dtor_ctors = dtor_ctor_thms,
+ ctor_dtors = ctor_dtor_thms, ctor_injects = ctor_inject_thms,
+ dtor_injects = dtor_inject_thms, xtor_maps = dtor_Jmap_thms, xtor_setss = dtor_Jset_thmss',
+ xtor_rels = dtor_Jrel_thms, xtor_un_fold_thms = dtor_unfold_thms,
+ xtor_co_rec_thms = dtor_corec_thms, xtor_co_rec_o_maps = dtor_corec_o_Jmap_thms,
+ xtor_rel_co_induct = Jrel_coinduct_thm, dtor_set_inducts = dtor_Jset_induct_thms,
xtor_co_rec_transfers = dtor_corec_transfer_thms};
in
timer; (fp_res, lthy')
--- a/src/HOL/Tools/BNF/bnf_lfp.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Tools/BNF/bnf_lfp.ML Mon Mar 30 11:36:30 2015 +0200
@@ -1826,13 +1826,14 @@
val lthy' = lthy |> note_all ? snd o Local_Theory.notes (common_notes @ notes @ Ibnf_notes);
val fp_res =
- {Ts = Ts, bnfs = Ibnfs, ctors = ctors, dtors = dtors, xtor_co_recs = recs,
- xtor_co_induct = ctor_induct_thm, dtor_ctors = dtor_ctor_thms, ctor_dtors = ctor_dtor_thms,
- ctor_injects = ctor_inject_thms, dtor_injects = dtor_inject_thms,
- xtor_maps = ctor_Imap_thms, xtor_setss = ctor_Iset_thmss',
- xtor_rels = ctor_Irel_thms, xtor_co_rec_thms = ctor_rec_thms,
- xtor_co_rec_o_maps = ctor_rec_o_Imap_thms, xtor_rel_co_induct = Irel_induct_thm,
- dtor_set_inducts = [], xtor_co_rec_transfers = ctor_rec_transfer_thms};
+ {Ts = Ts, bnfs = Ibnfs, ctors = ctors, dtors = dtors, xtor_un_folds = folds,
+ xtor_co_recs = recs, xtor_co_induct = ctor_induct_thm, dtor_ctors = dtor_ctor_thms,
+ ctor_dtors = ctor_dtor_thms, ctor_injects = ctor_inject_thms,
+ dtor_injects = dtor_inject_thms, xtor_maps = ctor_Imap_thms, xtor_setss = ctor_Iset_thmss',
+ xtor_rels = ctor_Irel_thms, xtor_un_fold_thms = ctor_fold_thms,
+ xtor_co_rec_thms = ctor_rec_thms, xtor_co_rec_o_maps = ctor_rec_o_Imap_thms,
+ xtor_rel_co_induct = Irel_induct_thm, dtor_set_inducts = [],
+ xtor_co_rec_transfers = ctor_rec_transfer_thms};
in
timer; (fp_res, lthy')
end;
--- a/src/HOL/Tools/BNF/bnf_lfp_basic_sugar.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Tools/BNF/bnf_lfp_basic_sugar.ML Mon Mar 30 11:36:30 2015 +0200
@@ -17,9 +17,7 @@
open BNF_FP_Def_Sugar
fun trivial_absT_info_of fpT =
- {absT = fpT,
- repT = fpT,
- abs = Const (@{const_name id_bnf}, fpT --> fpT),
+ {absT = fpT, repT = fpT, abs = Const (@{const_name id_bnf}, fpT --> fpT),
rep = Const (@{const_name id_bnf}, fpT --> fpT),
abs_inject = @{thm type_definition.Abs_inject[OF type_definition_id_bnf_UNIV UNIV_I UNIV_I]},
abs_inverse = @{thm type_definition.Abs_inverse[OF type_definition_id_bnf_UNIV UNIV_I]},
@@ -31,24 +29,19 @@
$> Morphism.term_morphism "BNF" (Term.map_types Logic.unvarifyT_global));
fun trivial_fp_result_of fp_bnf fpT C xtor_map xtor_sets xtor_rel ctor_rec_o_map xtor_rel_induct =
- {Ts = [fpT],
- bnfs = [fp_bnf],
- ctors = [Const (@{const_name xtor}, fpT --> fpT)],
- dtors = [Const (@{const_name xtor}, fpT --> fpT)],
- xtor_co_recs = [Const (@{const_name ctor_rec}, (fpT --> C) --> (fpT --> C))],
- xtor_co_induct = @{thm xtor_induct},
- dtor_ctors = @{thms xtor_xtor},
- ctor_dtors = @{thms xtor_xtor},
- ctor_injects = @{thms xtor_inject},
- dtor_injects = @{thms xtor_inject},
- xtor_maps = [xtor_map],
- xtor_setss = [xtor_sets],
- xtor_rels = [xtor_rel],
- xtor_co_rec_thms = [map_id0_of_bnf fp_bnf RS @{thm ctor_rec}],
- xtor_co_rec_o_maps = [ctor_rec_o_map],
- xtor_rel_co_induct = xtor_rel_induct,
- dtor_set_inducts = [],
- xtor_co_rec_transfers = []};
+ let
+ val xtors = [Const (@{const_name xtor}, fpT --> fpT)];
+ val co_recs = [Const (@{const_name ctor_rec}, (fpT --> C) --> (fpT --> C))];
+ val co_rec_thms = [map_id0_of_bnf fp_bnf RS @{thm ctor_rec}];
+ in
+ {Ts = [fpT], bnfs = [fp_bnf], ctors = xtors, dtors = xtors, xtor_un_folds = co_recs,
+ xtor_co_recs = co_recs, xtor_co_induct = @{thm xtor_induct}, dtor_ctors = @{thms xtor_xtor},
+ ctor_dtors = @{thms xtor_xtor}, ctor_injects = @{thms xtor_inject},
+ dtor_injects = @{thms xtor_inject}, xtor_maps = [xtor_map], xtor_setss = [xtor_sets],
+ xtor_rels = [xtor_rel], xtor_un_fold_thms = co_rec_thms, xtor_co_rec_thms = co_rec_thms,
+ xtor_co_rec_o_maps = [ctor_rec_o_map], xtor_rel_co_induct = xtor_rel_induct,
+ dtor_set_inducts = [], xtor_co_rec_transfers = []}
+ end;
fun fp_sugar_of_sum ctxt =
let
--- a/src/HOL/Tools/Lifting/lifting_bnf.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Tools/Lifting/lifting_bnf.ML Mon Mar 30 11:36:30 2015 +0200
@@ -93,7 +93,7 @@
in
[((Binding.empty, []), [([relator_eq_onp_thm], @{attributes [relator_eq_onp]})])]
end
- handle ERROR _ => [] (* the "lifting" plugin should work even if "quotient" is disabled *)
+ handle NO_PRED_DATA _ => [] (* the "lifting" plugin should work even if "quotient" is disabled *)
(* relator_mono *)
--- a/src/HOL/Tools/Lifting/lifting_term.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Tools/Lifting/lifting_term.ML Mon Mar 30 11:36:30 2015 +0200
@@ -128,7 +128,7 @@
Pretty.str "found."]))
end
-fun is_id_quot thm = (Thm.prop_of thm = Thm.prop_of @{thm identity_quotient}) (* FIXME equality!? *)
+fun is_id_quot thm = Thm.eq_thm_prop (thm, @{thm identity_quotient})
fun zip_Tvars ctxt type_name rty_Tvars qty_Tvars =
case try (get_rel_quot_thm ctxt) type_name of
--- a/src/HOL/Tools/Qelim/cooper.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Tools/Qelim/cooper.ML Mon Mar 30 11:36:30 2015 +0200
@@ -531,12 +531,12 @@
val tab = fold Termtab.update
(map (fn eq =>
let val (s,t) = Thm.cprop_of eq |> Thm.dest_arg |> Thm.dest_binop
- val th = if Thm.term_of s = Thm.term_of t (* FIXME equality? *)
- then the (Termtab.lookup inStab (Thm.term_of s))
- else FWD (instantiate' [] [SOME s, SOME t] eqelem_th)
- [eq, the (Termtab.lookup inStab (Thm.term_of s))]
- in (Thm.term_of t, th) end)
- sths) Termtab.empty
+ val th =
+ if s aconvc t
+ then the (Termtab.lookup inStab (Thm.term_of s))
+ else FWD (instantiate' [] [SOME s, SOME t] eqelem_th)
+ [eq, the (Termtab.lookup inStab (Thm.term_of s))]
+ in (Thm.term_of t, th) end) sths) Termtab.empty
in
fn ct => the (Termtab.lookup tab (Thm.term_of ct))
handle Option.Option =>
--- a/src/HOL/Tools/Quotient/quotient_term.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Tools/Quotient/quotient_term.ML Mon Mar 30 11:36:30 2015 +0200
@@ -355,7 +355,7 @@
| NONE => raise CODE_GEN ("get_relmap (no relation map function found for type " ^ s ^ ")"))
end
-fun is_id_quot thm = (Thm.prop_of thm = Thm.prop_of @{thm identity_quotient3}) (* FIXME equality *)
+fun is_id_quot thm = Thm.eq_thm_prop (thm, @{thm identity_quotient3})
open Lifting_Util
--- a/src/HOL/Tools/Transfer/transfer_bnf.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Tools/Transfer/transfer_bnf.ML Mon Mar 30 11:36:30 2015 +0200
@@ -6,6 +6,8 @@
signature TRANSFER_BNF =
sig
+ exception NO_PRED_DATA of unit
+
val transfer_plugin: string
val base_name_of_bnf: BNF_Def.bnf -> binding
val type_name_of_bnf: BNF_Def.bnf -> string
@@ -21,6 +23,8 @@
open BNF_FP_Util
open BNF_FP_Def_Sugar
+exception NO_PRED_DATA of unit
+
val transfer_plugin = Plugin_Name.declare_setup @{binding transfer}
(* util functions *)
@@ -318,9 +322,9 @@
(* simplification rules for the predicator *)
fun lookup_defined_pred_data lthy name =
- case (Transfer.lookup_pred_data lthy name) of
+ case Transfer.lookup_pred_data lthy name of
SOME data => data
- | NONE => (error "lookup_pred_data: something went utterly wrong")
+ | NONE => raise NO_PRED_DATA ()
fun lives_of_fp (fp_sugar: fp_sugar) =
let
@@ -377,6 +381,7 @@
|> Variable.export lthy old_lthy
|> map Drule.zero_var_indexes
end
+ handle NO_PRED_DATA () => []
(* fp_sugar interpretation *)
--- a/src/HOL/Tools/functor.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Tools/functor.ML Mon Mar 30 11:36:30 2015 +0200
@@ -41,8 +41,9 @@
(* type analysis *)
-fun term_with_typ ctxt T t = Envir.subst_term_types
- (Type.typ_match (Proof_Context.tsig_of ctxt) (fastype_of t, T) Vartab.empty) t;
+fun term_with_typ ctxt T t =
+ Envir.subst_term_types
+ (Sign.typ_match (Proof_Context.theory_of ctxt) (fastype_of t, T) Vartab.empty) t;
fun find_atomic ctxt T =
let
--- a/src/HOL/Tools/inductive.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Tools/inductive.ML Mon Mar 30 11:36:30 2015 +0200
@@ -38,6 +38,8 @@
(string * thm list) list * local_theory
val inductive_cases_i: (Attrib.binding * term list) list -> local_theory ->
(string * thm list) list * local_theory
+ val ind_cases_rules: Proof.context ->
+ string list -> (binding * string option * mixfix) list -> thm list
val inductive_simps: (Attrib.binding * string list) list -> local_theory ->
(string * thm list) list * local_theory
val inductive_simps_i: (Attrib.binding * term list) list -> local_theory ->
@@ -600,19 +602,22 @@
val inductive_cases = gen_inductive_cases Attrib.check_src Syntax.read_prop;
val inductive_cases_i = gen_inductive_cases (K I) Syntax.check_prop;
+
+(* ind_cases *)
+
+fun ind_cases_rules ctxt raw_props raw_fixes =
+ let
+ val (props, ctxt' ) = Specification.read_props raw_props raw_fixes ctxt;
+ val rules = Proof_Context.export ctxt' ctxt (map (mk_cases ctxt') props);
+ in rules end;
+
val _ =
Theory.setup
(Method.setup @{binding ind_cases}
- (Scan.lift (Scan.repeat1 Args.name_inner_syntax --
- Scan.optional (Args.$$$ "for" |-- Scan.repeat1 Args.binding) []) >>
- (fn (raw_props, fixes) => fn ctxt =>
- let
- val (_, ctxt') = Variable.add_fixes_binding fixes ctxt;
- val props = Syntax.read_props ctxt' raw_props;
- val ctxt'' = fold Variable.declare_term props ctxt';
- val rules = Proof_Context.export ctxt'' ctxt (map (mk_cases ctxt'') props)
- in Method.erule ctxt 0 rules end))
- "dynamic case analysis on predicates");
+ (Scan.lift (Scan.repeat1 Parse.prop -- Parse.for_fixes) >>
+ (fn (props, fixes) => fn ctxt =>
+ Method.erule ctxt 0 (ind_cases_rules ctxt props fixes)))
+ "case analysis for inductive definitions, based on simplified elimination rule");
(* derivation of simplified equation *)
--- a/src/HOL/UNITY/Comp.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/UNITY/Comp.thy Mon Mar 30 11:36:30 2015 +0200
@@ -186,7 +186,7 @@
apply simp
apply (subgoal_tac "G \<in> preserves (funPair v w) ")
prefer 2 apply simp
-apply (drule_tac P1 = "split ?Q" in preserves_subset_stable [THEN subsetD],
+apply (drule_tac P1 = "split Q" for Q in preserves_subset_stable [THEN subsetD],
auto)
done
@@ -198,8 +198,8 @@
(*The G case remains*)
apply (auto simp add: preserves_def stable_def constrains_def)
(*We have a G-action, so delete assumptions about F-actions*)
-apply (erule_tac V = "\<forall>act \<in> Acts F. ?P act" in thin_rl)
-apply (erule_tac V = "\<forall>z. \<forall>act \<in> Acts F. ?P z act" in thin_rl)
+apply (erule_tac V = "\<forall>act \<in> Acts F. P act" for P in thin_rl)
+apply (erule_tac V = "\<forall>z. \<forall>act \<in> Acts F. P z act" for P in thin_rl)
apply (subgoal_tac "v x = v xa")
apply auto
apply (erule order_trans, blast)
--- a/src/HOL/UNITY/Comp/Alloc.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/UNITY/Comp/Alloc.thy Mon Mar 30 11:36:30 2015 +0200
@@ -1080,7 +1080,7 @@
text{*Could go to Extend.ML*}
lemma bij_fst_inv_inv_eq: "bij f ==> fst (inv (%(x, u). inv f x) z) = f z"
apply (rule fst_inv_equalityI)
- apply (rule_tac f = "%z. (f z, ?h z) " in surjI)
+ apply (rule_tac f = "%z. (f z, h z)" for h in surjI)
apply (simp add: bij_is_inj inv_f_f)
apply (simp add: bij_is_surj surj_f_inv_f)
done
--- a/src/HOL/UNITY/ELT.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/UNITY/ELT.thy Mon Mar 30 11:36:30 2015 +0200
@@ -63,7 +63,7 @@
lemma givenBy_eq_all: "(givenBy v) = {A. ALL x:A. ALL y. v x = v y --> y: A}"
apply (unfold givenBy_def, safe)
-apply (rule_tac [2] x = "v ` ?u" in image_eqI, auto)
+apply (rule_tac [2] x = "v ` _" in image_eqI, auto)
done
lemma givenByI: "(!!x y. [| x:A; v x = v y |] ==> y: A) ==> A: givenBy v"
@@ -101,7 +101,7 @@
"[| A : givenBy v; B : givenBy v |] ==> A-B : givenBy v"
apply (simp (no_asm_use) add: givenBy_eq_Collect)
apply safe
-apply (rule_tac x = "%z. ?R z & ~ ?Q z" in exI)
+apply (rule_tac x = "%z. R z & ~ Q z" for R Q in exI)
unfolding set_diff_eq
apply auto
done
--- a/src/HOL/UNITY/Guar.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/UNITY/Guar.thy Mon Mar 30 11:36:30 2015 +0200
@@ -349,7 +349,7 @@
==> (\<forall>H. F ok H & G ok H & F\<squnion>H \<in> welldef & F\<squnion>H \<in> X --> G\<squnion>H \<in> X) =
(F \<in> welldef \<inter> X --> G \<in> X)"
apply (unfold strict_ex_prop_def, safe)
-apply (erule_tac x = SKIP and P = "%H. ?PP H --> ?RR H" in allE)
+apply (erule_tac x = SKIP and P = "%H. PP H --> RR H" for PP RR in allE)
apply (auto dest: Join_welldef_D1 Join_welldef_D2)
done
@@ -372,7 +372,7 @@
==> (\<forall>H. F ok H & G ok H & F\<squnion>H \<in> welldef & F\<squnion>H \<in> X --> G\<squnion>H \<in> X) =
(F \<in> welldef \<inter> X --> G \<in> X)"
apply (unfold strict_uv_prop_def, safe)
-apply (erule_tac x = SKIP and P = "%H. ?PP H --> ?RR H" in allE)
+apply (erule_tac x = SKIP and P = "%H. PP H --> RR H" for PP RR in allE)
apply (auto dest: Join_welldef_D1 Join_welldef_D2)
done
--- a/src/HOL/UNITY/ProgressSets.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/UNITY/ProgressSets.thy Mon Mar 30 11:36:30 2015 +0200
@@ -383,7 +383,7 @@
apply (drule equalityD1)
apply (rule subset_trans)
prefer 2 apply assumption
-apply (thin_tac "?U \<subseteq> X")
+apply (thin_tac "_ \<subseteq> X")
apply (cut_tac A=X in UN_singleton)
apply (erule subst)
apply (simp only: cl_UN lattice_latticeof
--- a/src/HOL/UNITY/Project.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/UNITY/Project.thy Mon Mar 30 11:36:30 2015 +0200
@@ -410,7 +410,7 @@
prefer 2
apply (simp add: stable_def constrains_def, blast)
(*back to main goal*)
-apply (erule_tac V = "?AA \<subseteq> -C \<union> ?BB" in thin_rl)
+apply (erule_tac V = "AA \<subseteq> -C \<union> BB" for AA BB in thin_rl)
apply (drule bspec, assumption)
apply (simp add: extend_set_def project_act_def, blast)
done
--- a/src/HOL/UNITY/Transformers.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/UNITY/Transformers.thy Mon Mar 30 11:36:30 2015 +0200
@@ -139,7 +139,7 @@
"[|T-B \<subseteq> awp F T; act \<in> Acts F|]
==> T \<inter> wens F act B \<subseteq> wens F act (T\<inter>B)"
apply (rule subset_wens)
-apply (rule_tac P="\<lambda>x. ?f x \<subseteq> ?b" in ssubst [OF wens_unfold])
+apply (rule_tac P="\<lambda>x. f x \<subseteq> b" for f b in ssubst [OF wens_unfold])
apply (simp add: wp_def awp_def, blast)
done
--- a/src/HOL/Wellfounded.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Wellfounded.thy Mon Mar 30 11:36:30 2015 +0200
@@ -232,7 +232,7 @@
prefer 2 apply blast
apply (rule_tac x = "{z. z:Q & (z,y) : r^*}" in allE)
apply assumption
-apply (erule_tac V = "ALL Q. (EX x. x : Q) --> ?P Q" in thin_rl)
+apply (erule_tac V = "ALL Q. (EX x. x : Q) --> P Q" for P in thin_rl)
--{*essential for speed*}
txt{*Blast with new substOccur fails*}
apply (fast intro: converse_rtrancl_into_rtrancl)
--- a/src/HOL/Word/Bit_Representation.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Word/Bit_Representation.thy Mon Mar 30 11:36:30 2015 +0200
@@ -224,7 +224,7 @@
apply (erule conjI)
apply (drule_tac x=0 in fun_cong, force)
apply (rule ext)
- apply (drule_tac x="Suc ?x" in fun_cong, force)
+ apply (drule_tac x="Suc x" for x in fun_cong, force)
done
show ?thesis
by (auto elim: bin_nth_lem)
--- a/src/HOL/Word/Misc_Typedef.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Word/Misc_Typedef.thy Mon Mar 30 11:36:30 2015 +0200
@@ -194,7 +194,7 @@
prefer 2
apply (simp add: comp_assoc)
apply (rule ext)
- apply (drule_tac f="?a o ?b" in fun_cong)
+ apply (drule_tac f="a o b" for a b in fun_cong)
apply simp
done
--- a/src/HOL/Word/Tools/word_lib.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Word/Tools/word_lib.ML Mon Mar 30 11:36:30 2015 +0200
@@ -1,11 +1,19 @@
(* Title: HOL/Word/Tools/word_lib.ML
Author: Sascha Boehme, TU Muenchen and Thomas Sewell, NICTA
-Helper routines for operating on the word type at the ML level.
+Helper routines for operating on the word type.
*)
+signature WORD_LIB =
+sig
+ val dest_binT: typ -> int
+ val is_wordT: typ -> bool
+ val dest_wordT: typ -> int
+ val mk_wordT: int -> typ
+end
-structure Word_Lib = struct
+structure Word_Lib: WORD_LIB =
+struct
fun dest_binT T =
(case T of
@@ -21,20 +29,19 @@
fun dest_wordT (Type (@{type_name word}, [T])) = dest_binT T
| dest_wordT T = raise TYPE ("dest_wordT", [T], [])
-local
- fun mk_bitT i T =
- if i = 0
- then Type (@{type_name "Numeral_Type.bit0"}, [T])
- else Type (@{type_name "Numeral_Type.bit1"}, [T])
+
+fun mk_bitT i T =
+ if i = 0
+ then Type (@{type_name "Numeral_Type.bit0"}, [T])
+ else Type (@{type_name "Numeral_Type.bit1"}, [T])
- fun mk_binT size =
- if size = 0 then @{typ "Numeral_Type.num0"}
- else if size = 1 then @{typ "Numeral_Type.num1"}
- else let val (q, r) = Integer.div_mod size 2 in mk_bitT r (mk_binT q) end
-in
+fun mk_binT size =
+ if size = 0 then @{typ "Numeral_Type.num0"}
+ else if size = 1 then @{typ "Numeral_Type.num1"}
+ else let val (q, r) = Integer.div_mod size 2 in mk_bitT r (mk_binT q) end
+
fun mk_wordT size =
if size >= 0 then Type (@{type_name "word"}, [mk_binT size])
- else raise TYPE ("mk_wordT: " ^ quote (string_of_int size), [], [])
-end
+ else raise TYPE ("mk_wordT: " ^ string_of_int size, [], [])
end
--- a/src/HOL/Word/Word.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/Word/Word.thy Mon Mar 30 11:36:30 2015 +0200
@@ -3190,7 +3190,7 @@
apply clarsimp
apply (case_tac x, force)
apply (case_tac m, auto)
- apply (drule_tac t="length ?xs" in sym)
+ apply (drule_tac t="length xs" for xs in sym)
apply (clarsimp simp: map2_def zip_replicate o_def)
done
@@ -3203,7 +3203,7 @@
apply clarsimp
apply (case_tac x, force)
apply (case_tac m, auto)
- apply (drule_tac t="length ?xs" in sym)
+ apply (drule_tac t="length xs" for xs in sym)
apply (clarsimp simp: map2_def zip_replicate o_def map_replicate_const)
done
@@ -3811,7 +3811,7 @@
lemma word_split_cat_alt:
"w = word_cat u v \<Longrightarrow> size u + size v <= size w \<Longrightarrow> word_split w = (u, v)"
- apply (case_tac "word_split ?w")
+ apply (case_tac "word_split w")
apply (rule trans, assumption)
apply (drule test_bit_split)
apply safe
@@ -4694,7 +4694,7 @@
apply simp
apply (case_tac "1 + (n - m) = 0")
apply (simp add: word_rec_0)
- apply (rule_tac f = "word_rec ?a ?b" in arg_cong)
+ apply (rule_tac f = "word_rec a b" for a b in arg_cong)
apply (rule_tac t="m" and s="m + (1 + (n - m))" in subst)
apply simp
apply (simp (no_asm_use))
--- a/src/HOL/ex/Cartouche_Examples.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/ex/Cartouche_Examples.thy Mon Mar 30 11:36:30 2015 +0200
@@ -144,7 +144,7 @@
setup -- "document antiquotation"
\<open>
Thy_Output.antiquotation @{binding ML_cartouche}
- (Scan.lift Args.cartouche_source_position) (fn {context, ...} => fn source =>
+ (Scan.lift Args.cartouche_input) (fn {context, ...} => fn source =>
let
val toks = ML_Lex.read "fn _ => (" @ ML_Lex.read_source false source @ ML_Lex.read ");";
val _ = ML_Context.eval_in (SOME context) ML_Compiler.flags (Input.pos_of source) toks;
@@ -179,7 +179,7 @@
ML \<open>
Outer_Syntax.command
@{command_spec "text_cartouche"} ""
- (Parse.opt_target -- Parse.source_position Parse.cartouche >> Thy_Output.document_command)
+ (Parse.opt_target -- Parse.input Parse.cartouche >> Thy_Output.document_command)
\<close>
text_cartouche
@@ -225,7 +225,7 @@
subsubsection \<open>Explicit version: method with cartouche argument\<close>
method_setup ml_tactic = \<open>
- Scan.lift Args.cartouche_source_position
+ Scan.lift Args.cartouche_input
>> (fn arg => fn ctxt => SIMPLE_METHOD (ML_Tactic.ml_tactic arg ctxt))
\<close>
@@ -246,7 +246,7 @@
subsubsection \<open>Implicit version: method with special name "cartouche" (dynamic!)\<close>
method_setup "cartouche" = \<open>
- Scan.lift Args.cartouche_source_position
+ Scan.lift Args.cartouche_input
>> (fn arg => fn ctxt => SIMPLE_METHOD (ML_Tactic.ml_tactic arg ctxt))
\<close>
--- a/src/HOL/ex/Dedekind_Real.thy Mon Mar 30 10:58:08 2015 +0200
+++ b/src/HOL/ex/Dedekind_Real.thy Mon Mar 30 11:36:30 2015 +0200
@@ -44,10 +44,16 @@
qed
-definition "preal = {A. cut A}"
+typedef preal = "Collect cut"
+ by (blast intro: cut_of_rat [OF zero_less_one])
-typedef preal = preal
- unfolding preal_def by (blast intro: cut_of_rat [OF zero_less_one])
+lemma Abs_preal_induct [induct type: preal]:
+ "(\<And>x. cut x \<Longrightarrow> P (Abs_preal x)) \<Longrightarrow> P x"
+ using Abs_preal_induct [of P x] by simp
+
+lemma Rep_preal:
+ "cut (Rep_preal x)"
+ using Rep_preal [of x] by simp
definition
psup :: "preal set => preal" where
@@ -111,34 +117,34 @@
declare Abs_preal_inject [simp]
declare Abs_preal_inverse [simp]
-lemma rat_mem_preal: "0 < q ==> {r::rat. 0 < r & r < q} \<in> preal"
-by (simp add: preal_def cut_of_rat)
+lemma rat_mem_preal: "0 < q ==> cut {r::rat. 0 < r & r < q}"
+by (simp add: cut_of_rat)
-lemma preal_nonempty: "A \<in> preal ==> \<exists>x\<in>A. 0 < x"
- unfolding preal_def cut_def [abs_def] by blast
+lemma preal_nonempty: "cut A ==> \<exists>x\<in>A. 0 < x"
+ unfolding cut_def [abs_def] by blast
-lemma preal_Ex_mem: "A \<in> preal \<Longrightarrow> \<exists>x. x \<in> A"
+lemma preal_Ex_mem: "cut A \<Longrightarrow> \<exists>x. x \<in> A"
apply (drule preal_nonempty)
apply fast
done
-lemma preal_imp_psubset_positives: "A \<in> preal ==> A < {r. 0 < r}"
- by (force simp add: preal_def cut_def)
+lemma preal_imp_psubset_positives: "cut A ==> A < {r. 0 < r}"
+ by (force simp add: cut_def)
-lemma preal_exists_bound: "A \<in> preal ==> \<exists>x. 0 < x & x \<notin> A"
+lemma preal_exists_bound: "cut A ==> \<exists>x. 0 < x & x \<notin> A"
apply (drule preal_imp_psubset_positives)
apply auto
done
-lemma preal_exists_greater: "[| A \<in> preal; y \<in> A |] ==> \<exists>u \<in> A. y < u"
- unfolding preal_def cut_def [abs_def] by blast
+lemma preal_exists_greater: "[| cut A; y \<in> A |] ==> \<exists>u \<in> A. y < u"
+ unfolding cut_def [abs_def] by blast
-lemma preal_downwards_closed: "[| A \<in> preal; y \<in> A; 0 < z; z < y |] ==> z \<in> A"
- unfolding preal_def cut_def [abs_def] by blast
+lemma preal_downwards_closed: "[| cut A; y \<in> A; 0 < z; z < y |] ==> z \<in> A"
+ unfolding cut_def [abs_def] by blast
text{*Relaxing the final premise*}
lemma preal_downwards_closed':
- "[| A \<in> preal; y \<in> A; 0 < z; z \<le> y |] ==> z \<in> A"
+ "[| cut A; y \<in> A; 0 < z; z \<le> y |] ==> z \<in> A"
apply (simp add: order_le_less)
apply (blast intro: preal_downwards_closed)
done
@@ -147,7 +153,7 @@
Gleason p. 122 - Remark (1)*}
lemma not_in_preal_ub:
- assumes A: "A \<in> preal"
+ assumes A: "cut A"
and notx: "x \<notin> A"
and y: "y \<in> A"
and pos: "0 < x"
@@ -167,6 +173,7 @@
text {* preal lemmas instantiated to @{term "Rep_preal X"} *}
lemma mem_Rep_preal_Ex: "\<exists>x. x \<in> Rep_preal X"
+thm preal_Ex_mem
by (rule preal_Ex_mem [OF Rep_preal])
lemma Rep_preal_exists_bound: "\<exists>x>0. x \<notin> Rep_preal X"
@@ -195,7 +202,7 @@
by (auto simp add: preal_le_def preal_less_def Rep_preal_inject)
qed
-lemma preal_imp_pos: "[|A \<in> preal; r \<in> A|] ==> 0 < r"
+lemma preal_imp_pos: "[|cut A; r \<in> A|] ==> 0 < r"
by (insert preal_imp_psubset_positives, blast)
instance preal :: linorder
@@ -237,7 +244,7 @@
text{*Part 1 of Dedekind sections definition*}
lemma add_set_not_empty:
- "[|A \<in> preal; B \<in> preal|] ==> {} \<subset> add_set A B"
+ "[|cut A; cut B|] ==> {} \<subset> add_set A B"
apply (drule preal_nonempty)+
apply (auto simp add: add_set_def)
done
@@ -245,7 +252,7 @@
text{*Part 2 of Dedekind sections definition. A structured version of
this proof is @{text preal_not_mem_mult_set_Ex} below.*}
lemma preal_not_mem_add_set_Ex:
- "[|A \<in> preal; B \<in> preal|] ==> \<exists>q>0. q \<notin> add_set A B"
+ "[|cut A; cut B|] ==> \<exists>q>0. q \<notin> add_set A B"
apply (insert preal_exists_bound [of A] preal_exists_bound [of B], auto)
apply (rule_tac x = "x+xa" in exI)
apply (simp add: add_set_def, clarify)
@@ -254,8 +261,8 @@
done
lemma add_set_not_rat_set:
- assumes A: "A \<in> preal"
- and B: "B \<in> preal"
+ assumes A: "cut A"
+ and B: "cut B"
shows "add_set A B < {r. 0 < r}"
proof
from preal_imp_pos [OF A] preal_imp_pos [OF B]
@@ -267,12 +274,12 @@
text{*Part 3 of Dedekind sections definition*}
lemma add_set_lemma3:
- "[|A \<in> preal; B \<in> preal; u \<in> add_set A B; 0 < z; z < u|]
+ "[|cut A; cut B; u \<in> add_set A B; 0 < z; z < u|]
==> z \<in> add_set A B"
proof (unfold add_set_def, clarify)
fix x::rat and y::rat
- assume A: "A \<in> preal"
- and B: "B \<in> preal"
+ assume A: "cut A"
+ and B: "cut B"
and [simp]: "0 < z"
and zless: "z < x + y"
and x: "x \<in> A"
@@ -310,7 +317,7 @@
text{*Part 4 of Dedekind sections definition*}
lemma add_set_lemma4:
- "[|A \<in> preal; B \<in> preal; y \<in> add_set A B|] ==> \<exists>u \<in> add_set A B. y < u"
+ "[|cut A; cut B; y \<in> add_set A B|] ==> \<exists>u \<in> add_set A B. y < u"
apply (auto simp add: add_set_def)
apply (frule preal_exists_greater [of A], auto)
apply (rule_tac x="u + ya" in exI)
@@ -318,8 +325,8 @@
done
lemma mem_add_set:
- "[|A \<in> preal; B \<in> preal|] ==> add_set A B \<in> preal"
-apply (simp (no_asm_simp) add: preal_def cut_def)
+ "[|cut A; cut B|] ==> cut (add_set A B)"
+apply (simp (no_asm_simp) add: cut_def)
apply (blast intro!: add_set_not_empty add_set_not_rat_set
add_set_lemma3 add_set_lemma4)
done
@@ -353,15 +360,15 @@
text{*Part 1 of Dedekind sections definition*}
lemma mult_set_not_empty:
- "[|A \<in> preal; B \<in> preal|] ==> {} \<subset> mult_set A B"
+ "[|cut A; cut B|] ==> {} \<subset> mult_set A B"
apply (insert preal_nonempty [of A] preal_nonempty [of B])
apply (auto simp add: mult_set_def)
done
text{*Part 2 of Dedekind sections definition*}
lemma preal_not_mem_mult_set_Ex:
- assumes A: "A \<in> preal"
- and B: "B \<in> preal"
+ assumes A: "cut A"
+ and B: "cut B"
shows "\<exists>q. 0 < q & q \<notin> mult_set A B"
proof -
from preal_exists_bound [OF A] obtain x where 1 [simp]: "0 < x" "x \<notin> A" by blast
@@ -389,8 +396,8 @@
qed
lemma mult_set_not_rat_set:
- assumes A: "A \<in> preal"
- and B: "B \<in> preal"
+ assumes A: "cut A"
+ and B: "cut B"
shows "mult_set A B < {r. 0 < r}"
proof
show "mult_set A B \<subseteq> {r. 0 < r}"
@@ -404,12 +411,12 @@
text{*Part 3 of Dedekind sections definition*}
lemma mult_set_lemma3:
- "[|A \<in> preal; B \<in> preal; u \<in> mult_set A B; 0 < z; z < u|]
+ "[|cut A; cut B; u \<in> mult_set A B; 0 < z; z < u|]
==> z \<in> mult_set A B"
proof (unfold mult_set_def, clarify)
fix x::rat and y::rat
- assume A: "A \<in> preal"
- and B: "B \<in> preal"
+ assume A: "cut A"
+ and B: "cut B"
and [simp]: "0 < z"
and zless: "z < x * y"
and x: "x \<in> A"
@@ -436,7 +443,7 @@
text{*Part 4 of Dedekind sections definition*}
lemma mult_set_lemma4:
- "[|A \<in> preal; B \<in> preal; y \<in> mult_set A B|] ==> \<exists>u \<in> mult_set A B. y < u"
+ "[|cut A; cut B; y \<in> mult_set A B|] ==> \<exists>u \<in> mult_set A B. y < u"
apply (auto simp add: mult_set_def)
apply (frule preal_exists_greater [of A], auto)
apply (rule_tac x="u * ya" in exI)
@@ -446,8 +453,8 @@
lemma mem_mult_set:
- "[|A \<in> preal; B \<in> preal|] ==> mult_set A B \<in> preal"
-apply (simp (no_asm_simp) add: preal_def cut_def)
+ "[|cut A; cut B|] ==> cut (mult_set A B)"
+apply (simp (no_asm_simp) add: cut_def)
apply (blast intro!: mult_set_not_empty mult_set_not_rat_set
mult_set_lemma3 mult_set_lemma4)
done
@@ -470,7 +477,7 @@
lemma preal_mult_1: "(1::preal) * z = z"
proof (induct z)
fix A :: "rat set"
- assume A: "A \<in> preal"
+ assume A: "cut A"
have "{w. \<exists>u. 0 < u \<and> u < 1 & (\<exists>v \<in> A. w = u * v)} = A" (is "?lhs = A")
proof
show "?lhs \<subseteq> A"
@@ -588,7 +595,7 @@
subsection{*Existence of Inverse, a Positive Real*}
lemma mem_inv_set_ex:
- assumes A: "A \<in> preal" shows "\<exists>x y. 0 < x & x < y & inverse y \<notin> A"
+ assumes A: "cut A" shows "\<exists>x y. 0 < x & x < y & inverse y \<notin> A"
proof -
from preal_exists_bound [OF A]
obtain x where [simp]: "0<x" "x \<notin> A" by blast
@@ -605,7 +612,7 @@
text{*Part 1 of Dedekind sections definition*}
lemma inverse_set_not_empty:
- "A \<in> preal ==> {} \<subset> inverse_set A"
+ "cut A ==> {} \<subset> inverse_set A"
apply (insert mem_inv_set_ex [of A])
apply (auto simp add: inverse_set_def)
done
@@ -613,7 +620,7 @@
text{*Part 2 of Dedekind sections definition*}
lemma preal_not_mem_inverse_set_Ex:
- assumes A: "A \<in> preal" shows "\<exists>q. 0 < q & q \<notin> inverse_set A"
+ assumes A: "cut A" shows "\<exists>q. 0 < q & q \<notin> inverse_set A"
proof -
from preal_nonempty [OF A]
obtain x where x: "x \<in> A" and xpos [simp]: "0<x" ..
@@ -635,7 +642,7 @@
qed
lemma inverse_set_not_rat_set:
- assumes A: "A \<in> preal" shows "inverse_set A < {r. 0 < r}"
+ assumes A: "cut A" shows "inverse_set A < {r. 0 < r}"
proof
show "inverse_set A \<subseteq> {r. 0 < r}" by (force simp add: inverse_set_def)
next
@@ -645,7 +652,7 @@
text{*Part 3 of Dedekind sections definition*}
lemma inverse_set_lemma3:
- "[|A \<in> preal; u \<in> inverse_set A; 0 < z; z < u|]
+ "[|cut A; u \<in> inverse_set A; 0 < z; z < u|]
==> z \<in> inverse_set A"
apply (auto simp add: inverse_set_def)
apply (auto intro: order_less_trans)
@@ -653,7 +660,7 @@
text{*Part 4 of Dedekind sections definition*}
lemma inverse_set_lemma4:
- "[|A \<in> preal; y \<in> inverse_set A|] ==> \<exists>u \<in> inverse_set A. y < u"
+ "[|cut A; y \<in> inverse_set A|] ==> \<exists>u \<in> inverse_set A. y < u"
apply (auto simp add: inverse_set_def)
apply (drule dense [of y])
apply (blast intro: order_less_trans)
@@ -661,8 +668,8 @@
lemma mem_inverse_set:
- "A \<in> preal ==> inverse_set A \<in> preal"
-apply (simp (no_asm_simp) add: preal_def cut_def)
+ "cut A ==> cut (inverse_set A)"
+apply (simp (no_asm_simp) add: cut_def)
apply (blast intro!: inverse_set_not_empty inverse_set_not_rat_set
inverse_set_lemma3 inverse_set_lemma4)
done
@@ -671,7 +678,7 @@
subsection{*Gleason's Lemma 9-3.4, page 122*}
lemma Gleason9_34_exists:
- assumes A: "A \<in> preal"
+ assumes A: "cut A"
and "\<forall>x\<in>A. x + u \<in> A"
and "0 \<le> z"
shows "\<exists>b\<in>A. b + (of_int z) * u \<in> A"
@@ -694,7 +701,7 @@
qed
lemma Gleason9_34_contra:
- assumes A: "A \<in> preal"
+ assumes A: "cut A"
shows "[|\<forall>x\<in>A. x + u \<in> A; 0 < u; 0 < y; y \<notin> A|] ==> False"
proof (induct u, induct y)
fix a::int and b::int
@@ -734,7 +741,7 @@
lemma Gleason9_34:
- assumes A: "A \<in> preal"
+ assumes A: "cut A"
and upos: "0 < u"
shows "\<exists>r \<in> A. r + u \<notin> A"
proof (rule ccontr, simp)
@@ -750,7 +757,7 @@
subsection{*Gleason's Lemma 9-3.6*}
lemma lemma_gleason9_36:
- assumes A: "A \<in> preal"
+ assumes A: "cut A"
and x: "1 < x"
shows "\<exists>r \<in> A. r*x \<notin> A"
proof -
@@ -965,8 +972,8 @@
done
lemma mem_diff_set:
- "R < S ==> diff_set (Rep_preal S) (Rep_preal R) \<in> preal"
-apply (unfold preal_def cut_def [abs_def])
+ "R < S ==> cut (diff_set (Rep_preal S) (Rep_preal R))"
+apply (unfold cut_def)
apply (blast intro!: diff_set_not_empty diff_set_not_rat_set
diff_set_lemma3 diff_set_lemma4)
done
@@ -1072,11 +1079,17 @@
lemma preal_add_left_less_cancel: "T + R < T + S ==> R < (S::preal)"
by (auto elim: preal_add_right_less_cancel simp add: preal_add_commute [of T])
-lemma preal_add_less_cancel_left: "(T + (R::preal) < T + S) = (R < S)"
+lemma preal_add_less_cancel_left [simp]: "(T + (R::preal) < T + S) = (R < S)"
by (blast intro: preal_add_less2_mono2 preal_add_left_less_cancel)
-lemma preal_add_le_cancel_left: "(T + (R::preal) \<le> T + S) = (R \<le> S)"
-by (simp add: linorder_not_less [symmetric] preal_add_less_cancel_left)
+lemma preal_add_less_cancel_right [simp]: "((R::preal) + T < S + T) = (R < S)"
+ using preal_add_less_cancel_left [symmetric, of R S T] by (simp add: ac_simps)
+
+lemma preal_add_le_cancel_left [simp]: "(T + (R::preal) \<le> T + S) = (R \<le> S)"
+by (simp add: linorder_not_less [symmetric])
+
+lemma preal_add_le_cancel_right [simp]: "((R::preal) + T \<le> S + T) = (R \<le> S)"
+ using preal_add_le_cancel_left [symmetric, of R S T] by (simp add: ac_simps)
lemma preal_add_right_cancel: "(R::preal) + T = S + T ==> R = S"
apply (insert linorder_less_linear [of R S], safe)
@@ -1086,10 +1099,9 @@
lemma preal_add_left_cancel: "C + A = C + B ==> A = (B::preal)"
by (auto intro: preal_add_right_cancel simp add: preal_add_commute)
-instance preal :: linordered_cancel_ab_semigroup_add
+instance preal :: linordered_ab_semigroup_add
proof
fix a b c :: preal
- show "a + b = a + c \<Longrightarrow> b = c" by (rule preal_add_left_cancel)
show "a \<le> b \<Longrightarrow> c + a \<le> c + b" by (simp only: preal_add_le_cancel_left)
qed
@@ -1134,8 +1146,8 @@
by (blast dest: Rep_preal [THEN preal_exists_greater])
lemma preal_sup:
- "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y|] ==> (\<Union>X \<in> P. Rep_preal(X)) \<in> preal"
-apply (unfold preal_def cut_def [abs_def])
+ "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y|] ==> cut (\<Union>X \<in> P. Rep_preal(X))"
+apply (unfold cut_def)
apply (blast intro!: preal_sup_set_not_empty preal_sup_set_not_rat_set
preal_sup_set_lemma3 preal_sup_set_lemma4)
done
@@ -1242,7 +1254,7 @@
also have "... = x2 + (x + y1)" by (simp add: assms)
also have "... = (x2 + y1) + x" by (simp add: ac_simps)
finally have "(x1 + y2) + x = (x2 + y1) + x" .
- thus ?thesis by (rule add_right_imp_eq)
+ thus ?thesis by (rule preal_add_right_cancel)
qed
@@ -1354,7 +1366,7 @@
UN_equiv_class2 [OF equiv_realrel equiv_realrel real_mult_congruent2])
lemma real_mult_commute: "(z::real) * w = w * z"
-by (cases z, cases w, simp add: real_mult ac_simps ac_simps)
+by (cases z, cases w, simp add: real_mult ac_simps)
lemma real_mult_assoc: "((z1::real) * z2) * z3 = z1 * (z2 * z3)"
apply (cases z1, cases z2, cases z3)
@@ -1376,8 +1388,8 @@
proof -
have "(1::preal) < 1 + 1"
by (simp add: preal_self_less_add_left)
- thus ?thesis
- by (simp add: real_zero_def real_one_def)
+ then show ?thesis
+ by (simp add: real_zero_def real_one_def neq_iff)
qed
instance real :: comm_ring_1
@@ -1444,7 +1456,7 @@
assumes eq: "a+b = c+d" and le: "c \<le> a"
shows "b \<le> (d::preal)"
proof -
- have "c+d \<le> a+d" by (simp add: le)
+ from le have "c+d \<le> a+d" by simp
hence "a+b \<le> a+d" by (simp add: eq)
thus "b \<le> d" by simp
qed
@@ -1614,11 +1626,9 @@
by (simp add: linorder_not_less [symmetric])
lemma real_of_preal_zero_less: "0 < real_of_preal m"
-apply (insert preal_self_less_add_left [of 1 m])
-apply (auto simp add: real_zero_def real_of_preal_def
- real_less_def real_le_def ac_simps)
-apply (rule_tac x="m + 1" in exI, rule_tac x="1" in exI)
-apply (simp add: ac_simps)
+using preal_self_less_add_left [of 1 m]
+apply (auto simp add: real_zero_def real_of_preal_def real_less_def real_le_def ac_simps neq_iff)
+apply (metis Rep_preal_self_subset add.commute preal_le_def)
done
lemma real_of_preal_minus_less_zero: "- real_of_preal m < 0"
--- a/src/Pure/General/completion.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/General/completion.ML Mon Mar 30 11:36:30 2015 +0200
@@ -9,6 +9,7 @@
type T
val names: Position.T -> (string * (string * string)) list -> T
val none: T
+ val make: string * Position.T -> ((string -> bool) -> (string * (string * string)) list) -> T
val reported_text: T -> string
val suppress_abbrevs: string -> Markup.T list
end;
@@ -16,12 +17,12 @@
structure Completion: COMPLETION =
struct
+(* completion of names *)
+
abstype T =
Completion of {pos: Position.T, total: int, names: (string * (string * string)) list}
with
-(* completion of names *)
-
fun dest (Completion args) = args;
fun names pos names =
@@ -34,6 +35,11 @@
val none = names Position.none [];
+fun make (name, pos) make_names =
+ if Position.is_reported pos andalso name <> "" andalso name <> "_"
+ then names pos (make_names (String.isPrefix (Name.clean name)))
+ else none;
+
fun reported_text completion =
let val {pos, total, names} = dest completion in
if Position.is_reported pos andalso not (null names) then
--- a/src/Pure/General/name_space.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/General/name_space.ML Mon Mar 30 11:36:30 2015 +0200
@@ -232,7 +232,7 @@
(* completion *)
fun completion context space (xname, pos) =
- if Position.is_reported pos andalso xname <> "" andalso xname <> "_" then
+ Completion.make (xname, pos) (fn completed =>
let
fun result_ord ((xname1, (_, name1)), (xname2, (_, name2))) =
(case bool_ord (apply2 Long_Name.is_local (name2, name1)) of
@@ -241,21 +241,18 @@
EQUAL => string_ord (xname1, xname2)
| ord => ord)
| ord => ord);
- val x = Name.clean xname;
val Name_Space {kind, internals, ...} = space;
val ext = extern_shortest (Context.proof_of context) space;
- val names =
- Change_Table.fold
- (fn (a, (name :: _, _)) =>
- if String.isPrefix x a andalso not (is_concealed space name)
- then
- let val a' = ext name
- in if a = a' then cons (a', (kind, name)) else I end
- else I
- | _ => I) internals []
- |> sort_distinct result_ord;
- in Completion.names pos names end
- else Completion.none;
+ in
+ Change_Table.fold
+ (fn (a, (name :: _, _)) =>
+ if completed a andalso not (is_concealed space name) then
+ let val a' = ext name
+ in if a = a' then cons (a', (kind, name)) else I end
+ else I
+ | _ => I) internals []
+ |> sort_distinct result_ord
+ end);
(* merge *)
--- a/src/Pure/Isar/args.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/Isar/args.ML Mon Mar 30 11:36:30 2015 +0200
@@ -22,11 +22,11 @@
val mode: string -> bool parser
val maybe: 'a parser -> 'a option parser
val cartouche_inner_syntax: string parser
- val cartouche_source_position: Input.source parser
- val text_source_position: Input.source parser
+ val cartouche_input: Input.source parser
+ val text_input: Input.source parser
val text: string parser
val name_inner_syntax: string parser
- val name_source_position: Input.source parser
+ val name_input: Input.source parser
val name: string parser
val binding: binding parser
val alt_name: string parser
@@ -108,14 +108,14 @@
val cartouche = Parse.token Parse.cartouche;
val cartouche_inner_syntax = cartouche >> Token.inner_syntax_of;
-val cartouche_source_position = cartouche >> Token.source_position_of;
+val cartouche_input = cartouche >> Token.input_of;
val text_token = named || Parse.token (Parse.verbatim || Parse.cartouche);
-val text_source_position = text_token >> Token.source_position_of;
+val text_input = text_token >> Token.input_of;
val text = text_token >> Token.content_of;
val name_inner_syntax = named >> Token.inner_syntax_of;
-val name_source_position = named >> Token.source_position_of;
+val name_input = named >> Token.input_of;
val name = named >> Token.content_of;
val binding = Parse.position name >> Binding.make;
@@ -158,12 +158,10 @@
named >> evaluate Token.Attribute (fn tok => att (Token.content_of tok, Token.pos_of tok));
fun text_declaration read =
- internal_declaration ||
- text_token >> evaluate Token.Declaration (read o Token.source_position_of);
+ internal_declaration || text_token >> evaluate Token.Declaration (read o Token.input_of);
fun cartouche_declaration read =
- internal_declaration ||
- cartouche >> evaluate Token.Declaration (read o Token.source_position_of);
+ internal_declaration || cartouche >> evaluate Token.Declaration (read o Token.input_of);
(* terms and types *)
--- a/src/Pure/Isar/element.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/Isar/element.ML Mon Mar 30 11:36:30 2015 +0200
@@ -236,7 +236,7 @@
val concl_term = Object_Logic.drop_judgment thy concl;
val fixes = fold_aterms (fn v as Free (x, T) =>
- if Variable.newly_fixed ctxt' ctxt x andalso not (v aconv concl_term)
+ if Variable.is_newly_fixed ctxt' ctxt x andalso not (v aconv concl_term)
then insert (op =) (Variable.revert_fixed ctxt' x, T) else I | _ => I) prop [] |> rev;
val (assumes, cases) = take_suffix (fn prem =>
is_elim andalso concl aconv Logic.strip_assums_concl prem) prems;
--- a/src/Pure/Isar/parse.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/Isar/parse.ML Mon Mar 30 11:36:30 2015 +0200
@@ -16,7 +16,7 @@
val token: 'a parser -> Token.T parser
val range: 'a parser -> ('a * Position.range) parser
val position: 'a parser -> ('a * Position.T) parser
- val source_position: 'a parser -> Input.source parser
+ val input: 'a parser -> Input.source parser
val inner_syntax: 'a parser -> string parser
val command_: string parser
val keyword: string parser
@@ -176,7 +176,7 @@
fun range scan = (Scan.ahead not_eof >> (Token.range_of o single)) -- scan >> Library.swap;
fun position scan = (Scan.ahead not_eof >> Token.pos_of) -- scan >> Library.swap;
-fun source_position atom = Scan.ahead atom |-- not_eof >> Token.source_position_of;
+fun input atom = Scan.ahead atom |-- not_eof >> Token.input_of;
fun inner_syntax atom = Scan.ahead atom |-- not_eof >> Token.inner_syntax_of;
fun kind k =
@@ -368,8 +368,8 @@
(* embedded source text *)
-val ML_source = source_position (group (fn () => "ML source") text);
-val document_source = source_position (group (fn () => "document source") text);
+val ML_source = input (group (fn () => "ML source") text);
+val document_source = input (group (fn () => "document source") text);
(* terms *)
--- a/src/Pure/Isar/proof_context.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/Isar/proof_context.ML Mon Mar 30 11:36:30 2015 +0200
@@ -26,6 +26,7 @@
val tsig_of: Proof.context -> Type.tsig
val set_defsort: sort -> Proof.context -> Proof.context
val default_sort: Proof.context -> indexname -> sort
+ val arity_sorts: Proof.context -> string -> sort -> sort list
val consts_of: Proof.context -> Consts.T
val set_syntax_mode: Syntax.mode -> Proof.context -> Proof.context
val restore_syntax_mode: Proof.context -> Proof.context -> Proof.context
@@ -289,6 +290,7 @@
val tsig_of = #1 o #tsig o rep_data;
val set_defsort = map_tsig o apfst o Type.set_defsort;
fun default_sort ctxt = the_default (Type.defaultS (tsig_of ctxt)) o Variable.def_sort ctxt;
+fun arity_sorts ctxt = Type.arity_sorts (Context.pretty ctxt) (tsig_of ctxt);
val consts_of = #1 o #consts o rep_data;
val cases_of = #cases o rep_data;
--- a/src/Pure/Isar/specification.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/Isar/specification.ML Mon Mar 30 11:36:30 2015 +0200
@@ -7,6 +7,10 @@
signature SPECIFICATION =
sig
+ val read_props: string list -> (binding * string option * mixfix) list -> Proof.context ->
+ term list * Proof.context
+ val read_prop: string -> (binding * string option * mixfix) list -> Proof.context ->
+ term * Proof.context
val check_spec:
(binding * typ option * mixfix) list -> (Attrib.binding * term) list -> Proof.context ->
(((binding * typ) * mixfix) list * (Attrib.binding * term) list) * Proof.context
@@ -79,6 +83,22 @@
structure Specification: SPECIFICATION =
struct
+(* prepare propositions *)
+
+fun read_props raw_props raw_fixes ctxt =
+ let
+ val ctxt1 = ctxt |> Proof_Context.read_vars raw_fixes |-> Proof_Context.add_fixes |> #2;
+ val props1 = map (Syntax.parse_prop ctxt1) raw_props;
+ val (props2, ctxt2) = ctxt1 |> fold_map Variable.fix_dummy_patterns props1;
+ val props3 = Syntax.check_props ctxt2 props2;
+ val ctxt3 = ctxt2 |> fold Variable.declare_term props3;
+ in (props3, ctxt3) end;
+
+fun read_prop raw_prop raw_fixes ctxt =
+ let val ([prop], ctxt') = read_props [raw_prop] raw_fixes ctxt
+ in (prop, ctxt') end;
+
+
(* prepare specification *)
local
--- a/src/Pure/Isar/token.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/Isar/token.ML Mon Mar 30 11:36:30 2015 +0200
@@ -56,7 +56,7 @@
val is_blank: T -> bool
val is_newline: T -> bool
val content_of: T -> string
- val source_position_of: T -> Input.source
+ val input_of: T -> Input.source
val inner_syntax_of: T -> string
val keyword_markup: bool * Markup.T -> string -> Markup.T
val completion_report: T -> Position.report_text list
@@ -281,12 +281,12 @@
fun content_of (Token (_, (_, x), _)) = x;
-fun source_position_of (Token ((source, range), (kind, _), _)) =
+fun input_of (Token ((source, range), (kind, _), _)) =
Input.source (delimited_kind kind) source range;
fun inner_syntax_of tok =
let val x = content_of tok
- in if YXML.detect x then x else Syntax.implode_input (source_position_of tok) end;
+ in if YXML.detect x then x else Syntax.implode_input (input_of tok) end;
(* markup reports *)
--- a/src/Pure/ML/ml_antiquotations.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/ML/ml_antiquotations.ML Mon Mar 30 11:36:30 2015 +0200
@@ -37,8 +37,19 @@
val _ = Theory.setup
(ML_Antiquotation.value @{binding system_option}
(Args.context -- Scan.lift (Parse.position Args.name) >> (fn (ctxt, (name, pos)) =>
- (Context_Position.report ctxt pos (Options.default_markup (name, pos));
- ML_Syntax.print_string name))) #>
+ let
+ val markup =
+ Options.default_markup (name, pos) handle ERROR msg =>
+ let
+ val completion =
+ Completion.make (name, pos) (fn completed =>
+ Options.names (Options.default ())
+ |> filter completed
+ |> map (fn a => (a, ("system_option", a))));
+ val report = Markup.markup_report (Completion.reported_text completion);
+ in cat_error msg report end;
+ val _ = Context_Position.report ctxt pos markup;
+ in ML_Syntax.print_string name end)) #>
ML_Antiquotation.value @{binding theory}
(Args.context -- Scan.lift (Parse.position Args.name) >> (fn (ctxt, (name, pos)) =>
--- a/src/Pure/PIDE/command.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/PIDE/command.ML Mon Mar 30 11:36:30 2015 +0200
@@ -181,7 +181,7 @@
Toplevel.setmp_thread_position tr
(fn () =>
Outer_Syntax.side_comments span |> maps (fn cmt =>
- (Thy_Output.check_text (Token.source_position_of cmt) st'; [])
+ (Thy_Output.check_text (Token.input_of cmt) st'; [])
handle exn =>
if Exn.is_interrupt exn then reraise exn
else Runtime.exn_messages_ids exn)) ();
--- a/src/Pure/Syntax/local_syntax.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/Syntax/local_syntax.ML Mon Mar 30 11:36:30 2015 +0200
@@ -55,7 +55,7 @@
val thy_syntax = Sign.syn_of thy;
fun update_gram ((true, add, m), decls) = Syntax.update_type_gram add m decls
| update_gram ((false, add, m), decls) =
- Syntax.update_const_gram add (Sign.is_logtype thy) m decls;
+ Syntax.update_const_gram add (Sign.logical_types thy) m decls;
val local_syntax = thy_syntax
|> fold update_gram (AList.coalesce (op =) (rev (map snd mixfixes)));
--- a/src/Pure/Syntax/mixfix.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/Syntax/mixfix.ML Mon Mar 30 11:36:30 2015 +0200
@@ -26,7 +26,7 @@
val make_type: int -> typ
val binder_name: string -> string
val syn_ext_types: (string * typ * mixfix) list -> Syntax_Ext.syn_ext
- val syn_ext_consts: (string -> bool) -> (string * typ * mixfix) list -> Syntax_Ext.syn_ext
+ val syn_ext_consts: string list -> (string * typ * mixfix) list -> Syntax_Ext.syn_ext
end;
structure Mixfix: MIXFIX =
@@ -119,7 +119,7 @@
val binder_stamp = stamp ();
val binder_name = suffix "_binder";
-fun syn_ext_consts is_logtype const_decls =
+fun syn_ext_consts logical_types const_decls =
let
fun mk_infix sy ty c p1 p2 p3 =
[Syntax_Ext.Mfix ("op " ^ sy, ty, c, [], 1000),
@@ -152,7 +152,7 @@
val consts = binders @ map (fn (c, _, _) => (c, "")) const_decls;
in
- Syntax_Ext.syn_ext' is_logtype mfix consts ([], binder_trs, binder_trs', []) ([], [])
+ Syntax_Ext.syn_ext' logical_types mfix consts ([], binder_trs, binder_trs', []) ([], [])
end;
end;
--- a/src/Pure/Syntax/syntax.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/Syntax/syntax.ML Mon Mar 30 11:36:30 2015 +0200
@@ -108,7 +108,7 @@
(string * ((Proof.context -> typ -> term list -> term) * stamp)) list *
(string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list -> syntax -> syntax
val update_type_gram: bool -> mode -> (string * typ * mixfix) list -> syntax -> syntax
- val update_const_gram: bool -> (string -> bool) ->
+ val update_const_gram: bool -> string list ->
mode -> (string * typ * mixfix) list -> syntax -> syntax
val update_trrules: Ast.ast trrule list -> syntax -> syntax
val remove_trrules: Ast.ast trrule list -> syntax -> syntax
@@ -659,8 +659,8 @@
fun update_type_gram add prmode decls =
(if add then update_syntax else remove_syntax) prmode (Mixfix.syn_ext_types decls);
-fun update_const_gram add is_logtype prmode decls =
- (if add then update_syntax else remove_syntax) prmode (Mixfix.syn_ext_consts is_logtype decls);
+fun update_const_gram add logical_types prmode decls =
+ (if add then update_syntax else remove_syntax) prmode (Mixfix.syn_ext_consts logical_types decls);
val update_trrules = update_syntax mode_default o Syntax_Ext.syn_ext_rules o check_rules;
val remove_trrules = remove_syntax mode_default o Syntax_Ext.syn_ext_rules o check_rules;
--- a/src/Pure/Syntax/syntax_ext.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/Syntax/syntax_ext.ML Mon Mar 30 11:36:30 2015 +0200
@@ -31,7 +31,7 @@
val mfix_delims: string -> string list
val mfix_args: string -> int
val escape: string -> string
- val syn_ext': (string -> bool) -> mfix list ->
+ val syn_ext': string list -> mfix list ->
(string * string) list -> (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list *
(string * ((Proof.context -> term list -> term) * stamp)) list *
(string * ((Proof.context -> typ -> term list -> term) * stamp)) list *
@@ -184,8 +184,10 @@
(* mfix_to_xprod *)
-fun mfix_to_xprod is_logtype (mfix as Mfix (sy, typ, const, pris, pri)) =
+fun mfix_to_xprod logical_types (mfix as Mfix (sy, typ, const, pris, pri)) =
let
+ val is_logtype = member (op =) logical_types;
+
fun check_pri p =
if p >= 0 andalso p <= 1000 then ()
else err_in_mfix ("Precedence out of range: " ^ string_of_int p) mfix;
@@ -288,12 +290,12 @@
(* syn_ext *)
-fun syn_ext' is_logtype mfixes consts trfuns (parse_rules, print_rules) =
+fun syn_ext' logical_types mfixes consts trfuns (parse_rules, print_rules) =
let
val (parse_ast_translation, parse_translation, print_translation,
print_ast_translation) = trfuns;
- val xprod_results = map (mfix_to_xprod is_logtype) mfixes;
+ val xprod_results = map (mfix_to_xprod logical_types) mfixes;
val xprods = map #1 xprod_results;
val consts' = map_filter #2 xprod_results;
val parse_rules' = rev (map_filter #3 xprod_results);
@@ -311,7 +313,7 @@
end;
-val syn_ext = syn_ext' (K false);
+val syn_ext = syn_ext' [];
fun syn_ext_rules rules = syn_ext [] [] ([], [], [], []) rules;
fun syn_ext_trfuns trfuns = syn_ext [] [] trfuns ([], []);
--- a/src/Pure/System/options.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/System/options.ML Mon Mar 30 11:36:30 2015 +0200
@@ -13,6 +13,7 @@
val unknownT: string
type T
val empty: T
+ val names: T -> string list
val markup: T -> string * Position.T -> Markup.T
val typ: T -> string -> string
val bool: T -> string -> bool
@@ -59,6 +60,8 @@
val empty = Options Symtab.empty;
+fun names (Options tab) = sort_strings (Symtab.keys tab);
+
(* check *)
--- a/src/Pure/System/options.scala Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/System/options.scala Mon Mar 30 11:36:30 2015 +0200
@@ -93,9 +93,9 @@
{
command(SECTION) ~! text ^^
{ case _ ~ a => (options: Options) => options.set_section(a) } |
- opt(command(PUBLIC)) ~ position(command(OPTION)) ~! (option_name ~ $$$(":") ~ option_type ~
+ opt(command(PUBLIC)) ~ command(OPTION) ~! (position(option_name) ~ $$$(":") ~ option_type ~
$$$("=") ~ option_value ~ ($$$("--") ~! text ^^ { case _ ~ x => x } | success(""))) ^^
- { case a ~ ((_, pos)) ~ (b ~ _ ~ c ~ _ ~ d ~ e) =>
+ { case a ~ _ ~ ((b, pos) ~ _ ~ c ~ _ ~ d ~ e) =>
(options: Options) => options.declare(a.isDefined, pos, b, c, d, e) }
}
--- a/src/Pure/Thy/latex.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/Thy/latex.ML Mon Mar 30 11:36:30 2015 +0200
@@ -132,7 +132,7 @@
enclose "{\\isacharbackquoteopen}" "{\\isacharbackquoteclose}" (output_syms s)
else if Token.is_kind Token.Verbatim tok then
let
- val ants = Antiquote.read (Token.source_position_of tok);
+ val ants = Antiquote.read (Token.input_of tok);
val out = implode (map output_syms_antiq ants);
in enclose "{\\isacharverbatimopen}" "{\\isacharverbatimclose}" out end
else if Token.is_kind Token.Cartouche tok then
--- a/src/Pure/Thy/thy_output.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/Thy/thy_output.ML Mon Mar 30 11:36:30 2015 +0200
@@ -600,7 +600,7 @@
basic_entity @{binding typ} Args.typ_abbrev Syntax.pretty_typ #>
basic_entity @{binding class} (Scan.lift Args.name_inner_syntax) pretty_class #>
basic_entity @{binding type} (Scan.lift Args.name) pretty_type #>
- basic_entity @{binding text} (Scan.lift Args.text_source_position) pretty_text_report #>
+ basic_entity @{binding text} (Scan.lift Args.text_input) pretty_text_report #>
basic_entities @{binding prf} Attrib.thms (pretty_prf false) #>
basic_entities @{binding full_prf} Attrib.thms (pretty_prf true) #>
basic_entity @{binding theory} (Scan.lift (Parse.position Args.name)) pretty_theory);
@@ -672,7 +672,7 @@
local
-fun ml_text name ml = antiquotation name (Scan.lift Args.text_source_position)
+fun ml_text name ml = antiquotation name (Scan.lift Args.text_input)
(fn {context = ctxt, ...} => fn source =>
(ML_Context.eval_in (SOME ctxt) ML_Compiler.flags (Input.pos_of source) (ml source);
verbatim_text ctxt (Input.source_content source)));
--- a/src/Pure/Thy/thy_syntax.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/Thy/thy_syntax.ML Mon Mar 30 11:36:30 2015 +0200
@@ -27,7 +27,7 @@
fun reports_of_token keywords tok =
let
val malformed_symbols =
- Input.source_explode (Token.source_position_of tok)
+ Input.source_explode (Token.input_of tok)
|> map_filter (fn (sym, pos) =>
if Symbol.is_malformed sym
then SOME ((pos, Markup.bad), "Malformed symbolic character") else NONE);
--- a/src/Pure/Tools/build.scala Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/Tools/build.scala Mon Mar 30 11:36:30 2015 +0200
@@ -242,8 +242,8 @@
{ case _ ~ (_ ~ (x ~ _)) => x } | success("document")) ~
rep1(path)) ^^ { case _ ~ (x ~ y) => y.map((x, _)) }
- position(command(SESSION)) ~!
- (session_name ~
+ command(SESSION) ~!
+ (position(session_name) ~
(($$$("(") ~! (rep1(name) <~ $$$(")")) ^^ { case _ ~ x => x }) | success(Nil)) ~
(($$$(IN) ~! path ^^ { case _ ~ x => x }) | success(".")) ~
($$$("=") ~!
@@ -253,7 +253,7 @@
rep1(theories) ~
(($$$(FILES) ~! rep1(path) ^^ { case _ ~ x => x }) | success(Nil)) ~
(rep(document_files) ^^ (x => x.flatten))))) ^^
- { case (_, pos) ~ (a ~ b ~ c ~ (_ ~ (d ~ e ~ f ~ g ~ h ~ i))) =>
+ { case _ ~ ((a, pos) ~ b ~ c ~ (_ ~ (d ~ e ~ f ~ g ~ h ~ i))) =>
Session_Entry(pos, a, b, c, d, e, f, g, h, i) }
}
--- a/src/Pure/Tools/rail.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/Tools/rail.ML Mon Mar 30 11:36:30 2015 +0200
@@ -364,7 +364,7 @@
val _ = Theory.setup
(Thy_Output.antiquotation @{binding rail}
- (Scan.lift (Parse.source_position (Parse.string || Parse.cartouche)))
+ (Scan.lift (Parse.input (Parse.string || Parse.cartouche)))
(fn {state, context, ...} => output_rules state o read context));
end;
--- a/src/Pure/Tools/rule_insts.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/Tools/rule_insts.ML Mon Mar 30 11:36:30 2015 +0200
@@ -13,7 +13,9 @@
(binding * string option * mixfix) list -> thm -> thm
val read_instantiate: Proof.context ->
((indexname * Position.T) * string) list -> string list -> thm -> thm
+ val read_term: string -> Proof.context -> term * Proof.context
val schematic: bool Config.T
+ val goal_context: term -> Proof.context -> (string * typ) list * Proof.context
val res_inst_tac: Proof.context ->
((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm ->
int -> tactic
@@ -69,19 +71,6 @@
else error_var "Bad sort for instantiation of type variable: " (xi, pos)
end;
-fun read_terms ss Ts ctxt =
- let
- fun parse T = if T = propT then Syntax.parse_prop ctxt else Syntax.parse_term ctxt;
- val (ts, ctxt') = fold_map Variable.fix_dummy_patterns (map2 parse Ts ss) ctxt;
- val ts' =
- map2 (Type.constraint o Type_Infer.paramify_vars) Ts ts
- |> Syntax.check_terms ctxt'
- |> Variable.polymorphic ctxt';
- val Ts' = map Term.fastype_of ts';
- val tyenv = fold Type.raw_match (Ts ~~ Ts') Vartab.empty;
- val tyenv' = Vartab.fold (fn (xi, (S, T)) => cons ((xi, S), T)) tyenv [];
- in ((ts', tyenv'), ctxt') end;
-
fun make_instT f v =
let
val T = TVar v;
@@ -94,24 +83,45 @@
val t' = f t;
in if t aconv t' then NONE else SOME (v, t') end;
+fun read_terms ss Ts ctxt =
+ let
+ fun parse T = if T = propT then Syntax.parse_prop ctxt else Syntax.parse_term ctxt;
+ val (ts, ctxt') = fold_map Variable.fix_dummy_patterns (map2 parse Ts ss) ctxt;
+ val ts' =
+ map2 (Type.constraint o Type_Infer.paramify_vars) Ts ts
+ |> Syntax.check_terms ctxt'
+ |> Variable.polymorphic ctxt';
+ val Ts' = map Term.fastype_of ts';
+ val tyenv = fold (Sign.typ_match (Proof_Context.theory_of ctxt)) (Ts ~~ Ts') Vartab.empty;
+ val tyenv' = Vartab.fold (fn (xi, (S, T)) => cons ((xi, S), T)) tyenv [];
+ in ((ts', tyenv'), ctxt') end;
+
in
+fun read_term s ctxt =
+ let
+ val (t, ctxt') = Variable.fix_dummy_patterns (Syntax.parse_term ctxt s) ctxt;
+ val t' = Syntax.check_term ctxt' t;
+ in (t', ctxt') end;
+
fun read_insts thm mixed_insts ctxt =
let
val (type_insts, term_insts) =
List.partition (fn (((x, _), _), _) => String.isPrefix "'" x) mixed_insts;
+ (*thm context*)
+ val ctxt1 = Variable.declare_thm thm ctxt;
val tvars = Thm.fold_terms Term.add_tvars thm [];
val vars = Thm.fold_terms Term.add_vars thm [];
(*explicit type instantiations*)
- val instT1 = Term_Subst.instantiateT (map (read_type ctxt tvars) type_insts);
+ val instT1 = Term_Subst.instantiateT (map (read_type ctxt1 tvars) type_insts);
val vars1 = map (apsnd instT1) vars;
(*term instantiations*)
val (xs, ss) = split_list term_insts;
val Ts = map (the_type vars1) xs;
- val ((ts, inferred), ctxt') = read_terms ss Ts ctxt;
+ val ((ts, inferred), ctxt2) = read_terms ss Ts ctxt1;
(*implicit type instantiations*)
val instT2 = Term_Subst.instantiateT inferred;
@@ -122,7 +132,7 @@
val inst_tvars = map_filter (make_instT (instT2 o instT1)) tvars;
val inst_vars = map_filter (make_inst inst2) vars2;
- in ((inst_tvars, inst_vars), ctxt') end;
+ in ((inst_tvars, inst_vars), ctxt2) end;
end;
@@ -132,9 +142,7 @@
fun where_rule ctxt mixed_insts fixes thm =
let
- val ctxt' = ctxt
- |> Variable.declare_thm thm
- |> Proof_Context.read_vars fixes |-> Proof_Context.add_fixes |> #2;
+ val ctxt' = ctxt |> Proof_Context.read_vars fixes |-> Proof_Context.add_fixes |> #2;
val ((inst_tvars, inst_vars), ctxt'') = read_insts thm mixed_insts ctxt';
in
thm
@@ -199,59 +207,57 @@
(** tactics **)
-val schematic = Attrib.setup_config_bool @{binding rule_insts_schematic} (K true);
+(* goal context *)
+
+(*legacy*)
+val schematic = Attrib.setup_config_bool @{binding rule_insts_schematic} (K false);
+
+fun goal_context goal ctxt =
+ let
+ val ((_, params), ctxt') = ctxt
+ |> Variable.declare_constraints goal
+ |> Variable.improper_fixes
+ |> Variable.focus_params goal
+ ||> Variable.restore_proper_fixes ctxt
+ ||> Config.get ctxt schematic ? Proof_Context.set_mode Proof_Context.mode_schematic;
+ in (params, ctxt') end;
+
(* resolution after lifting and instantiation; may refer to parameters of the subgoal *)
fun bires_inst_tac bires_flag ctxt mixed_insts fixes thm i st = CSUBGOAL (fn (cgoal, _) =>
let
- (* goal parameters *)
+ (* goal context *)
- val goal = Thm.term_of cgoal;
- val params =
- Logic.strip_params goal
- (*as they are printed: bound variables with the same name are renamed*)
- |> Term.rename_wrt_term goal
- |> rev;
- val (param_names, param_ctxt) = ctxt
- |> Variable.declare_thm thm
- |> Thm.fold_terms Variable.declare_constraints st
- |> Variable.improper_fixes
- |> Proof_Context.add_fixes (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) params)
- ||> Variable.restore_proper_fixes ctxt
- ||> Config.get ctxt schematic ? Proof_Context.set_mode Proof_Context.mode_schematic;
+ val (params, goal_ctxt) = goal_context (Thm.term_of cgoal) ctxt;
val paramTs = map #2 params;
- (* local fixes *)
+ (* instantiation context *)
- val (fixed, fixes_ctxt) = param_ctxt
- |> Proof_Context.read_vars fixes |-> Proof_Context.add_fixes;
+ val ((inst_tvars, inst_vars), inst_ctxt) = goal_ctxt
+ |> Proof_Context.read_vars fixes |-> Proof_Context.add_fixes |> #2
+ |> read_insts thm mixed_insts;
- val ((inst_tvars, inst_vars), inst_ctxt) = read_insts thm mixed_insts fixes_ctxt;
+ val fixed = map #1 (fold (Variable.add_newly_fixed inst_ctxt goal_ctxt o #2) inst_vars []);
(* lift and instantiate rule *)
val inc = Thm.maxidx_of st + 1;
- fun lift_var ((a, j), T) =
- Var ((a, j + inc), paramTs ---> Logic.incr_tvar inc T);
- fun lift_term t =
- fold_rev Term.absfree (param_names ~~ paramTs)
- (Logic.incr_indexes (fixed, paramTs, inc) t);
+ val lift_type = Logic.incr_tvar inc;
+ fun lift_var ((a, j), T) = Var ((a, j + inc), paramTs ---> lift_type T);
+ fun lift_term t = fold_rev Term.absfree params (Logic.incr_indexes (fixed, paramTs, inc) t);
val inst_tvars' = inst_tvars
- |> map (apply2 (Thm.ctyp_of inst_ctxt o Logic.incr_tvar inc) o apfst TVar);
+ |> map (apply2 (Thm.ctyp_of inst_ctxt o lift_type) o apfst TVar);
val inst_vars' = inst_vars
|> map (fn (v, t) => apply2 (Thm.cterm_of inst_ctxt) (lift_var v, lift_term t));
- val thm' =
- Drule.instantiate_normalize (inst_tvars', inst_vars')
- (Thm.lift_rule cgoal thm)
- |> singleton (Variable.export inst_ctxt param_ctxt);
- in
- compose_tac param_ctxt (bires_flag, thm', Thm.nprems_of thm) i
- end) i st;
+ val thm' = Thm.lift_rule cgoal thm
+ |> Drule.instantiate_normalize (inst_tvars', inst_vars')
+ |> singleton (Variable.export inst_ctxt ctxt);
+ in compose_tac ctxt (bires_flag, thm', Thm.nprems_of thm) i end) i st;
val res_inst_tac = bires_inst_tac false;
val eres_inst_tac = bires_inst_tac true;
--- a/src/Pure/sign.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/sign.ML Mon Mar 30 11:36:30 2015 +0200
@@ -24,7 +24,7 @@
val of_sort: theory -> typ * sort -> bool
val inter_sort: theory -> sort * sort -> sort
val witness_sorts: theory -> (typ * sort) list -> sort list -> (typ * sort) list
- val is_logtype: theory -> string -> bool
+ val logical_types: theory -> string list
val typ_instance: theory -> typ * typ -> bool
val typ_equiv: theory -> typ * typ -> bool
val typ_match: theory -> typ * typ -> Type.tyenv -> Type.tyenv
@@ -190,7 +190,7 @@
val of_sort = Type.of_sort o tsig_of;
val inter_sort = Type.inter_sort o tsig_of;
val witness_sorts = Type.witness_sorts o tsig_of;
-val is_logtype = member (op =) o Type.logical_types o tsig_of;
+val logical_types = Type.logical_types o tsig_of;
val typ_instance = Type.typ_instance o tsig_of;
fun typ_equiv thy (T, U) = typ_instance thy (T, U) andalso typ_instance thy (U, T);
@@ -369,7 +369,7 @@
val ctxt = Type.set_mode Type.mode_syntax (Proof_Context.init_global thy);
fun prep (c, T, mx) = (c, certify_typ_mode Type.mode_syntax thy (parse_typ ctxt T), mx)
handle ERROR msg => cat_error msg ("in syntax declaration " ^ quote c);
- in thy |> map_syn (change_gram (is_logtype thy) mode (map prep args)) end;
+ in thy |> map_syn (change_gram (logical_types thy) mode (map prep args)) end;
fun gen_add_syntax x = gen_syntax (Syntax.update_const_gram true) x;
--- a/src/Pure/type.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/type.ML Mon Mar 30 11:36:30 2015 +0200
@@ -49,7 +49,6 @@
val restore_mode: Proof.context -> Proof.context -> Proof.context
val type_space: tsig -> Name_Space.T
val type_alias: Name_Space.naming -> binding -> string -> tsig -> tsig
- val is_logtype: tsig -> string -> bool
val check_decl: Context.generic -> tsig ->
xstring * Position.T -> (string * Position.report list) * decl
val the_decl: tsig -> string * Position.T -> decl
@@ -252,8 +251,6 @@
fun type_alias naming binding name = map_tsig (fn (classes, default, types) =>
(classes, default, (Name_Space.alias_table naming binding name types)));
-val is_logtype = member (op =) o logical_types;
-
fun undecl_type c = "Undeclared type constructor: " ^ quote c;
--- a/src/Pure/type_infer_context.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/type_infer_context.ML Mon Mar 30 11:36:30 2015 +0200
@@ -169,7 +169,7 @@
fun unify ctxt =
let
val thy = Proof_Context.theory_of ctxt;
- val arity_sorts = Type.arity_sorts (Context.pretty ctxt) (Sign.tsig_of thy);
+ val arity_sorts = Proof_Context.arity_sorts ctxt;
(* adjust sorts of parameters *)
--- a/src/Pure/variable.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Pure/variable.ML Mon Mar 30 11:36:30 2015 +0200
@@ -38,7 +38,7 @@
val restore_proper_fixes: Proof.context -> Proof.context -> Proof.context
val is_improper: Proof.context -> string -> bool
val is_fixed: Proof.context -> string -> bool
- val newly_fixed: Proof.context -> Proof.context -> string -> bool
+ val is_newly_fixed: Proof.context -> Proof.context -> string -> bool
val fixed_ord: Proof.context -> string * string -> order
val intern_fixed: Proof.context -> string -> string
val markup_fixed: Proof.context -> string -> Markup.T
@@ -46,6 +46,8 @@
val revert_fixed: Proof.context -> string -> string
val add_fixed_names: Proof.context -> term -> string list -> string list
val add_fixed: Proof.context -> term -> (string * typ) list -> (string * typ) list
+ val add_newly_fixed: Proof.context -> Proof.context ->
+ term -> (string * typ) list -> (string * typ) list
val add_free_names: Proof.context -> term -> string list -> string list
val add_frees: Proof.context -> term -> (string * typ) list -> (string * typ) list
val add_fixes_binding: binding list -> Proof.context -> string list * Proof.context
@@ -73,6 +75,7 @@
(((ctyp * ctyp) list * (cterm * cterm) list) * thm list) * Proof.context
val tradeT: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list
val trade: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list
+ val focus_params: term -> Proof.context -> (string list * (string * typ) list) * Proof.context
val focus: term -> Proof.context -> ((string * (string * typ)) list * term) * Proof.context
val focus_cterm: cterm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context
val focus_subgoal: int -> thm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context
@@ -345,7 +348,7 @@
(* specialized name space *)
val is_fixed = Name_Space.defined_entry o fixes_space;
-fun newly_fixed inner outer = is_fixed inner andf (not o is_fixed outer);
+fun is_newly_fixed inner outer = is_fixed inner andf (not o is_fixed outer);
val fixed_ord = Name_Space.entry_ord o fixes_space;
val intern_fixed = Name_Space.intern o fixes_space;
@@ -382,6 +385,9 @@
fun add_fixed ctxt =
fold_aterms (fn Free (x, T) => is_fixed ctxt x ? insert (op =) (x, T) | _ => I);
+fun add_newly_fixed ctxt' ctxt =
+ fold_aterms (fn Free (x, T) => is_newly_fixed ctxt' ctxt x ? insert (op =) (x, T) | _ => I);
+
(* declarations *)
@@ -473,7 +479,7 @@
fun export_inst inner outer =
let
val declared_outer = is_declared outer;
- val still_fixed = not o newly_fixed inner outer;
+ val still_fixed = not o is_newly_fixed inner outer;
val gen_fixes =
Name_Space.fold_table (fn (y, _) => not (is_fixed outer y) ? cons y)
--- a/src/Tools/Code/code_preproc.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Tools/Code/code_preproc.ML Mon Mar 30 11:36:30 2015 +0200
@@ -371,8 +371,9 @@
val thy = Proof_Context.theory_of ctxt;
val all_classes = complete_proper_sort thy [class];
val super_classes = remove (op =) class all_classes;
- val classess = map (complete_proper_sort thy)
- (Sign.arity_sorts thy tyco [class]);
+ val classess =
+ map (complete_proper_sort thy)
+ (Proof_Context.arity_sorts ctxt tyco [class]);
val inst_params = inst_params thy tyco all_classes;
in (classess, (super_classes, inst_params)) end;
@@ -410,7 +411,7 @@
|> apfst (Vargraph.add_edge (c_k, c_k'))
end
and ensure_typmatch_inst ctxt arities eqngr (tyco, styps) class vardeps_data =
- if can (Sign.arity_sorts (Proof_Context.theory_of ctxt) tyco) [class]
+ if can (Proof_Context.arity_sorts ctxt tyco) [class]
then vardeps_data
|> ensure_inst ctxt arities eqngr (class, tyco)
|> fold_index (fn (k, styp) =>
--- a/src/Tools/induct.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Tools/induct.ML Mon Mar 30 11:36:30 2015 +0200
@@ -394,12 +394,11 @@
fun prep_inst ctxt align tune (tm, ts) =
let
- val cert = Thm.global_cterm_of (Proof_Context.theory_of ctxt);
fun prep_var (x, SOME t) =
let
- val cx = cert x;
+ val cx = Thm.cterm_of ctxt x;
val xT = Thm.typ_of_cterm cx;
- val ct = cert (tune t);
+ val ct = Thm.cterm_of ctxt (tune t);
val tT = Thm.typ_of_cterm ct;
in
if Type.could_unify (tT, xT) then SOME (cx, ct)
@@ -478,8 +477,6 @@
fun cases_tac ctxt simp insts opt_rule facts =
let
- val thy = Proof_Context.theory_of ctxt;
-
fun inst_rule r =
(if null insts then r
else
@@ -505,7 +502,7 @@
val rule' = rule
|> simp ? (simplified_rule' ctxt #> unmark_constraints ctxt);
in
- CASES (Rule_Cases.make_common (thy,
+ CASES (Rule_Cases.make_common (Proof_Context.theory_of ctxt,
Thm.prop_of (Rule_Cases.internalize_params rule')) cases)
((Method.insert_tac more_facts THEN' resolve_tac ctxt [rule'] THEN_ALL_NEW
(if simp then TRY o trivial_tac ctxt else K all_tac)) i) st
@@ -569,21 +566,18 @@
local
-fun dest_env thy env =
+fun dest_env ctxt env =
let
- val cert = Thm.global_cterm_of thy;
- val certT = Thm.global_ctyp_of thy;
val pairs = Vartab.dest (Envir.term_env env);
val types = Vartab.dest (Envir.type_env env);
- val ts = map (cert o Envir.norm_term env o #2 o #2) pairs;
- val xs = map2 (curry (cert o Var)) (map #1 pairs) (map Thm.typ_of_cterm ts);
- in (map (fn (xi, (S, T)) => (certT (TVar (xi, S)), certT T)) types, xs ~~ ts) end;
+ val ts = map (Thm.cterm_of ctxt o Envir.norm_term env o #2 o #2) pairs;
+ val xs = map2 (curry (Thm.cterm_of ctxt o Var)) (map #1 pairs) (map Thm.typ_of_cterm ts);
+ in (map (fn (xi, (S, T)) => apply2 (Thm.ctyp_of ctxt) (TVar (xi, S), T)) types, xs ~~ ts) end;
in
fun guess_instance ctxt rule i st =
let
- val thy = Proof_Context.theory_of ctxt;
val maxidx = Thm.maxidx_of st;
val goal = Thm.term_of (Thm.cprem_of st i); (*exception Subscript*)
val params = rev (Term.rename_wrt_term goal (Logic.strip_params goal));
@@ -599,7 +593,7 @@
in
Unify.smash_unifiers (Context.Proof ctxt)
[(Thm.concl_of rule', concl)] (Envir.empty (Thm.maxidx_of rule'))
- |> Seq.map (fn env => Drule.instantiate_normalize (dest_env thy env) rule')
+ |> Seq.map (fn env => Drule.instantiate_normalize (dest_env ctxt env) rule')
end
end
handle General.Subscript => Seq.empty;
@@ -654,19 +648,17 @@
fun meta_spec_tac ctxt n (x, T) = SUBGOAL (fn (goal, i) =>
let
- val thy = Proof_Context.theory_of ctxt;
- val cert = Thm.global_cterm_of thy;
-
val v = Free (x, T);
fun spec_rule prfx (xs, body) =
@{thm Pure.meta_spec}
|> Thm.rename_params_rule ([Name.clean (Variable.revert_fixed ctxt x)], 1)
- |> Thm.lift_rule (cert prfx)
+ |> Thm.lift_rule (Thm.cterm_of ctxt prfx)
|> `(Thm.prop_of #> Logic.strip_assums_concl)
|-> (fn pred $ arg =>
Drule.cterm_instantiate
- [(cert (Term.head_of pred), cert (Logic.rlist_abs (xs, body))),
- (cert (Term.head_of arg), cert (Logic.rlist_abs (xs, v)))]);
+ (map (apply2 (Thm.cterm_of ctxt))
+ [(Term.head_of pred, Logic.rlist_abs (xs, body)),
+ (Term.head_of arg, Logic.rlist_abs (xs, v))]));
fun goal_concl k xs (Const (@{const_name Pure.all}, _) $ Abs (a, T, B)) =
goal_concl k ((a, T) :: xs) B
@@ -835,8 +827,6 @@
fun coinduct_tac ctxt inst taking opt_rule facts = SUBGOAL_CASES (fn (goal, i, st) =>
let
- val thy = Proof_Context.theory_of ctxt;
-
fun inst_rule r =
if null inst then `Rule_Cases.get r
else Drule.cterm_instantiate (prep_inst ctxt align_right I (main_prop_of r, inst)) r
@@ -857,7 +847,7 @@
guess_instance ctxt rule i st
|> Seq.map (rule_instance ctxt (burrow_options (Variable.polymorphic ctxt) taking))
|> Seq.maps (fn rule' =>
- CASES (Rule_Cases.make_common (thy,
+ CASES (Rule_Cases.make_common (Proof_Context.theory_of ctxt,
Thm.prop_of (Rule_Cases.internalize_params rule')) cases)
(Method.insert_tac more_facts i THEN resolve_tac ctxt [rule'] i) st))
end);
--- a/src/Tools/induct_tacs.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Tools/induct_tacs.ML Mon Mar 30 11:36:30 2015 +0200
@@ -7,11 +7,9 @@
signature INDUCT_TACS =
sig
val case_tac: Proof.context -> string -> (binding * string option * mixfix) list ->
- int -> tactic
- val case_rule_tac: Proof.context -> string -> (binding * string option * mixfix) list ->
- thm -> int -> tactic
- val induct_tac: Proof.context -> string option list list -> int -> tactic
- val induct_rules_tac: Proof.context -> string option list list -> thm list -> int -> tactic
+ thm option -> int -> tactic
+ val induct_tac: Proof.context -> string option list list ->
+ thm list option -> int -> tactic
end
structure Induct_Tacs: INDUCT_TACS =
@@ -28,35 +26,29 @@
quote (Syntax.string_of_term ctxt u) ^ Position.here pos);
in (u, U) end;
-fun gen_case_tac ctxt s fixes opt_rule i st =
+fun case_tac ctxt s fixes opt_rule = SUBGOAL (fn (goal, i) =>
let
val rule =
(case opt_rule of
SOME rule => rule
| NONE =>
let
- val ctxt' = ctxt
- |> Variable.focus_subgoal i st |> #2
- |> Config.get ctxt Rule_Insts.schematic ?
- Proof_Context.set_mode Proof_Context.mode_schematic
- |> Context_Position.set_visible false;
- val t = Syntax.read_term ctxt' s;
+ val (t, ctxt') = ctxt
+ |> Rule_Insts.goal_context goal |> #2
+ |> Context_Position.set_visible false
+ |> Rule_Insts.read_term s;
+ val pos = Syntax.read_input_pos s;
in
- (case Induct.find_casesT ctxt' (#2 (check_type ctxt' (t, Syntax.read_input_pos s))) of
+ (case Induct.find_casesT ctxt' (#2 (check_type ctxt' (t, pos))) of
rule :: _ => rule
| [] => @{thm case_split})
end);
val _ = Method.trace ctxt [rule];
-
- val xi =
- (case Induct.vars_of (Thm.term_of (Thm.cprem_of rule 1)) of
- Var (xi, _) :: _ => xi
- | _ => raise THM ("Malformed cases rule", 0, [rule]));
- in Rule_Insts.res_inst_tac ctxt [((xi, Position.none), s)] fixes rule i st end
- handle THM _ => Seq.empty;
-
-fun case_tac ctxt s xs = gen_case_tac ctxt s xs NONE;
-fun case_rule_tac ctxt s xs rule = gen_case_tac ctxt s xs (SOME rule);
+ in
+ (case Induct.vars_of (Thm.term_of (Thm.cprem_of rule 1)) handle THM _ => [] of
+ Var (xi, _) :: _ => Rule_Insts.res_inst_tac ctxt [((xi, Position.none), s)] fixes rule i
+ | _ => no_tac)
+ end);
(* induction *)
@@ -72,32 +64,34 @@
in
-fun gen_induct_tac ctxt0 varss opt_rules i st =
+fun induct_tac ctxt varss opt_rules i st =
let
- val ((_, goal), ctxt) = Variable.focus_subgoal i st ctxt0;
- val (prems, concl) = Logic.strip_horn (Thm.term_of goal);
+ val goal = Thm.cprem_of st i;
+ val (_, goal_ctxt) = Rule_Insts.goal_context (Thm.term_of goal) ctxt
+ and ((_, goal'), _) = Variable.focus_cterm goal ctxt;
+ val (prems, concl) = Logic.strip_horn (Thm.term_of goal');
fun induct_var name =
let
- val t = Syntax.read_term ctxt name;
+ val t = Syntax.read_term goal_ctxt name;
val pos = Syntax.read_input_pos name;
val (x, _) = Term.dest_Free t handle TERM _ =>
error ("Induction argument not a variable: " ^
- quote (Syntax.string_of_term ctxt t) ^ Position.here pos);
+ quote (Syntax.string_of_term goal_ctxt t) ^ Position.here pos);
val eq_x = fn Free (y, _) => x = y | _ => false;
val _ =
if Term.exists_subterm eq_x concl then ()
else
error ("No such variable in subgoal: " ^
- quote (Syntax.string_of_term ctxt t) ^ Position.here pos);
+ quote (Syntax.string_of_term goal_ctxt t) ^ Position.here pos);
val _ =
if (exists o Term.exists_subterm) eq_x prems then
warning ("Induction variable occurs also among premises: " ^
- quote (Syntax.string_of_term ctxt t) ^ Position.here pos)
+ quote (Syntax.string_of_term goal_ctxt t) ^ Position.here pos)
else ();
- in #1 (check_type ctxt (t, pos)) end;
+ in #1 (check_type goal_ctxt (t, pos)) end;
- val argss = map (map (Option.map induct_var)) varss;
+ val argss = (map o map o Option.map) induct_var varss;
val rule =
(case opt_rules of
SOME rules => #2 (Rule_Cases.strict_mutual_rule ctxt rules)
@@ -117,9 +111,6 @@
end;
-fun induct_tac ctxt args = gen_induct_tac ctxt args NONE;
-fun induct_rules_tac ctxt args rules = gen_induct_tac ctxt args (SOME rules);
-
(* method syntax *)
@@ -139,14 +130,13 @@
Theory.setup
(Method.setup @{binding case_tac}
(Args.goal_spec -- Scan.lift (Args.name_inner_syntax -- Parse.for_fixes) -- opt_rule >>
- (fn ((quant, (s, xs)), r) => fn ctxt => SIMPLE_METHOD'' quant (gen_case_tac ctxt s xs r)))
+ (fn ((quant, (s, xs)), r) => fn ctxt => SIMPLE_METHOD'' quant (case_tac ctxt s xs r)))
"unstructured case analysis on types" #>
Method.setup @{binding induct_tac}
(Args.goal_spec -- varss -- opt_rules >>
- (fn ((quant, vs), rs) => fn ctxt => SIMPLE_METHOD'' quant (gen_induct_tac ctxt vs rs)))
+ (fn ((quant, vs), rs) => fn ctxt => SIMPLE_METHOD'' quant (induct_tac ctxt vs rs)))
"unstructured induction on types");
end;
end;
-
--- a/src/Tools/subtyping.ML Mon Mar 30 10:58:08 2015 +0200
+++ b/src/Tools/subtyping.ML Mon Mar 30 11:36:30 2015 +0200
@@ -144,7 +144,7 @@
fun unify weak ctxt =
let
val thy = Proof_Context.theory_of ctxt;
- val arity_sorts = Type.arity_sorts (Context.pretty ctxt) (Sign.tsig_of thy);
+ val arity_sorts = Proof_Context.arity_sorts ctxt;
(* adjust sorts of parameters *)
@@ -344,7 +344,7 @@
val coes_graph = coes_graph_of ctxt;
val tmaps = tmaps_of ctxt;
- val arity_sorts = Type.arity_sorts (Context.pretty ctxt) (Sign.tsig_of thy);
+ val arity_sorts = Proof_Context.arity_sorts ctxt;
fun split_cs _ [] = ([], [])
| split_cs f (c :: cs) =