I am attempting to use the fundamental theorem of calculus to prove the lemma lm1:
lemma lm1:
fixes f :: "real ⇒ real"
assumes "∀x∈{a..b}. (f has_vector_derivative f' x) (at x within {a .. b})"
"∀x∈{a..b}. a ≤ x"
"∀x∈{a..b}. x ≤ b"
shows "∀x∈{a..b}. (f' has_integral (f x - f a)) {a .. x}"
This is just an extension of the following lemma lm2, which proves the integration over the whole set.
lemma lm2:
fixes f :: "real ⇒ real"
assumes "∀x∈{a..b}. (f has_vector_derivative f' x) (at x within {a .. b})"
"a ≤ b"
shows "∀x∈{a..b}. (f' has_integral (f b - f a)) {a .. b}"
using assms
apply(simp add: fundamental_theorem_of_calculus)
Instead, I want to prove that the integration is true for any value in the set, not just the lower and upper bound. How am I able to show this?
first of all, the second and third assumption in lm1 are trivial:
lemma "∀x∈{a..b}. a ≤ x ⟷ True" by simp
lemma "∀x∈{a..b}. x ≤ b ⟷ True" by simp
Therefore, you better assume "a <= b". To apply the fundamental_theorem_of_calculus "for any value in the set", you also need the derivative for variable bounds,
∀x∈{a..b}. ∀y∈{a..x}. (f has_vector_derivative f' y) (at y within {a..x})
for which you can use has_vector_derivative_within_subset in your proof:
lemma lm1':
fixes f :: "real ⇒ real"
assumes "∀x∈{a..b}. (f has_vector_derivative f' x) (at x within {a .. b})"
assumes "a ≤ b"
shows "∀x∈{a..b}. (f' has_integral (f x - f a)) {a .. x}"
using assms
apply safe
apply (rule fundamental_theorem_of_calculus)
apply simp
apply safe
apply (rule has_vector_derivative_within_subset[where s="{a .. b}"])
apply simp
apply simp
done
or more compact:
using assms
by (auto intro!: fundamental_theorem_of_calculus
intro: has_vector_derivative_within_subset[where s="{a .. b}"])
Related
I'm having some issues trying to do exercise 4.5 of 'Concrete Semantics' in Isar:
inductive S :: "alpha list ⇒ bool" where
Sε : "S []" |
aSb : "S m ⟹ S (a#m # [b])" |
SS : "S l ⟹ S r ⟹ S (l # r)"
inductive T :: "alpha list ⇒ bool" where
Tε : "T []" |
TaTb : "T l ⟹ T r ⟹ T (l # a#(r # [b]))"
lemma TS: "T w ⟹ S w"
proof (induction w rule: T.induct)
case Tε
show ?case by (simp add:Sε)
case (TaTb l r) show ?case using TaTb.IH(1) (* This being S l, which allows us to case-split on l using S.induct *)
proof (cases "l" rule: S.induct)
case Sε
then show ?case by (simp add: TaTb.IH(2) aSb)
next case (aSb m)
I'm getting Illegal schematic variable(s) in case "aSb"⌂
Also I find suspicious that in Sε I cannot refer to ?case, I get Unbound schematic variable: ?case. I'm thinking that maybe the problem is that I have a cases in an induction?
As summarized by the comments, you have two problems:
cases "l" rule: S.induct makes little sense and you should either use a nested induction induction l rule: S.induct or a case distinction cases l rule: S.cases
In cases you should use ?thesis instead of cases as the Isabelle/jEdit outline tells you (you can click on that thing to insert it into the buffer!). That way you would also have given a name to all variable in the case TaTb.
So you probably want something like:
lemma TS: "T w ⟹ S w"
proof (induction w rule: T.induct)
case Tε
show ?case by (simp add:Sε)
next
case (TaTb l r a b) show ?case using TaTb.IH(1)
proof (cases "l" rule: S.cases)
case Sε
then show ?thesis sorry
next
case (aSb m a b)
then show ?thesis sorry
next
case (SS l r)
then show ?thesis sorry
qed
qed
I am trying to proof a property of the icmp6 checksum function (sum 16bit integers, add carry, invert 16bit integer).
I defined the functions in isabelle. (I know my proofs are terrible)
But for some reason, isabelle can't proof something about the icmp_csum function, it wants to have.
When I replace the oops in the paste with done it produces thousands of lines that just says:
"linarith_split_limit exceeded (current value is 9)"
theory Scratch
imports Main Int List
begin
fun norm_helper :: "nat ⇒ nat" where
"norm_helper x = (let y = divide x 65536 in (y + x - y * 65536))"
lemma "x ≥ 65536 ⟹ norm_helper x < x" by simp
lemma h: "norm_helper x ≤ x" by simp
fun normalize :: "nat ⇒ nat" where
"normalize x = (if x ≥ 65536
then normalize (norm_helper x)
else x)"
inductive norm_to :: "nat ⇒ nat ⇒ bool" where
"(x < 65536) ⟹ norm_to x x"
| "norm_to y z ⟹ y = norm_helper x ⟹ norm_to x z"
lemma ne: "norm_to x y ⟹ y = normalize x"
apply (induct x y rule: norm_to.induct) by simp+
lemma i: "norm_to x y ⟹ x ≥ y"
apply (induct x y rule: norm_to.induct) by simp+
lemma l: "norm_to x y ⟹ y < 65536"
apply (induct x y rule: norm_to.induct) by simp+
lemma en: "y = normalize x ⟹ norm_to x y"
apply (induct x rule: normalize.induct)
proof -
fix x :: nat
assume 1: "(x ≥ 65536 ⟹ y = Scratch.normalize (norm_helper x) ⟹ norm_to (norm_helper x) y)"
assume 2: "y = Scratch.normalize x"
show "norm_to x y"
proof (cases "x ≥ 65536")
show "¬ 65536 ≤ x ⟹ norm_to x y"
using norm_to.intros(1)[of x] 2 by simp
{
assume s: "65536 ≤ x"
have d: "y = normalize (norm_helper x)" using 2 s by simp
show "65536 ≤ x ⟹ norm_to x y"
using 1 d norm_to.intros(2)[of "norm_helper x" y x]
by blast
}
qed
qed
lemma "normalize x ≤ x" using en i by simp
lemma n[simp]: "normalize x < 65536" using en l by blast
fun sum :: "nat list ⇒ nat" where
"sum [] = 0"
| "sum (x#xs) = x + sum xs"
fun csum :: "nat list ⇒ nat" where
"csum xs = normalize (sum xs)"
fun invert :: "nat ⇒ nat" where
"invert x = 65535 - x"
lemma c: "csum xs ≤ 65535" using n[of "sum xs"] by simp
lemma ic: "invert (csum xs) ≥ 0" using c[of xs] by blast
lemma asdf:
assumes "xs = ys"
shows "invert (csum xs) = invert (csum ys)"
using HOL.arg_cong[of "csum xs" "csum ys" invert,
OF HOL.arg_cong[of xs ys csum]] assms(1)
by blast
function icmp_csum :: "nat list ⇒ nat" where
"icmp_csum xs = invert (csum xs)"
apply simp
apply (rule asdf)
apply simp
oops
end
I have no idea why there is tracing output from linarith there, but given that your definition is neither recursive nor performs pattern matching, you can write it as a definition:
definition icmp_csum :: "nat list ⇒ nat" where
"icmp_csum xs = invert (csum xs)"
Another possibility is to change invert to a definition instead of a fun. (In general, if it's neither recursive nor performs pattern matching, definition is preferable because it has much less overhead than fun.)
NB, just import Main, not Main Int List.
Edit: An explanation from Tobias Nipkow on the mailing list:
This is a known issue. In the outdated LNCS 2283 you can find a discussion what to do about it in Section 3.5.3 Simplification and Recursive Functions. The gist: don't use "if", use pattern matching or "case". Or disable if_split.
I've written some simple parser combinators (without backtracking etc.). Here are the important definitions for my problem.
type_synonym ('a, 's) parser = "'s list ⇒ ('a * 's list) option"
definition sequenceP :: "('a, 's) parser
⇒ ('b, 's) parser
⇒ ('b, 's) parser" (infixl ">>P" 60) where
"sequenceP p q ≡ λ i .
(case p i of
None ⇒ None
| Some v ⇒ q (snd v))"
definition consumerP :: "('a, 's) parser ⇒ bool" where
"consumerP p ≡ (∀ i . (case p i of
None ⇒ True |
Some v ⇒ length (snd v) ≤ length i))"
I do want to proof the following lemma.
lemma consumerPI: "consumerP p ⟹ consumerP q ⟹ consumerP (p >>P q)"
apply (unfold sequenceP_def)
apply (simp (no_asm) add:consumerP_def)
apply clarsimp
apply (case_tac "case p i of None ⇒ None | Some v ⇒ q (snd v)")
apply simp
apply clarsimp
apply (case_tac "p i")
apply simp
apply clarsimp
apply (unfold consumerP_def)
I arrive at this proof state, at which I fail to continue.
goal (1 subgoal):
1. ⋀i a b aa ba.
⟦∀i. case p i of None ⇒ True | Some v ⇒ length (snd v) ≤ length i;
∀i. case q i of None ⇒ True | Some v ⇒ length (snd v) ≤ length i; q ba = Some (a, b); p i = Some (aa, ba)⟧
⟹ length b ≤ length i
Can anybody give me a tip how to solve this goal?
Thanks in advance!
It turns out that if you just want to prove the lemma, without further insight, then
lemma consumerPI: "consumerP p ⟹ consumerP q ⟹ consumerP (p >>P q)"
by (smt consumerP_def le_trans option.case_eq_if sequenceP_def)
does the job.
If you want to have insight, you want to go for a structured proof. First identify some useful lemmas about consumerP, and then write a Isar proof that details the necessary steps.
lemma consumerPI[intro!]:
assumes "⋀ i x r . p i = Some (x,r) ⟹ length r ≤ length i"
shows "consumerP p"
unfolding consumerP_def by (auto split: option.split elim: assms)
lemma consumerPE[elim, consumes 1]:
assumes "consumerP p"
assumes "p i = Some (x,r)"
shows "length r ≤ length i"
using assms by (auto simp add: consumerP_def split: option.split_asm)
lemma consumerP_sequencePI: "consumerP p ⟹ consumerP q ⟹ consumerP (p >>P q)"
proof-
assume "consumerP p"
assume "consumerP q"
show "consumerP (p >>P q)"
proof(rule consumerPI)
fix i x r
assume "(p >>P q) i = Some (x, r)"
then obtain x' r' where "p i = Some (x', r')" and "q r' = Some (x,r)"
by (auto simp add: sequenceP_def split:option.split_asm)
from `consumerP q` and `q r' = Some (x, r)`
have "length r ≤ length r'" by (rule consumerPE)
also
from `consumerP p` and `p i = Some (x', r')`
have "length r' ≤ length i" by (rule consumerPE)
finally
show "length r ≤ length i".
qed
qed
In fact, for this definition you can very nicely use the inductive command, and get intro and elim rules for free:
inductive consumerP where
consumerPI: "(⋀ i x r . p i = Some (x,r) ⟹ length r ≤ length i) ⟹ consumerP p"
In the above proof, you can replace by (rule consumerPE) by by cases and it works.
I want to (miss-)use Isabelle to show that two given formulas are syntactically equivalent. For example A ∧ B = B ∧ A.
I don't want to go into any detail with regards towards the logic behind the formulas. I don't want to care that A ∧ B is true when both A and B are true. I just want to compare the two formulas on a structural level and say that they are equivalent because of the commutative property.
Basicly i want to be abled to write lemmas comparing 2 formulas with some equality function and use the given, however they are to be specified, axioms.
So far i thought that this could and should be done using axiomatization, but everyone here tells me axiomatzation is bad.
This leads me to my question how should this task be done. How can 2, say propositional, formulas be compared in Isabelle with regards towards their syntactical equivalence. To give a concrete example:
formula ∧ formula | formula ∨ formula
are given as operators, if possible as datatype
and distributive and commutative property are given as rules.
A ∧ B = B ∧ A, if stated in a theorem should be provable.
That's what i want to do, i hope the idea is clear and someone can explain to me how to pursue this properly in Isabelle. Thanks in advance.
From what you've written, I'm pretty sure you mean syntactic equivalence. Two formulas are semantically equivalent if they evaluate to the same result for all valuations of variables; two formulas are syntactically equivalent if you can rewrite one to another given a certain set of rewrite rules (or, more generally, prove their equivalence using a certain set of inference rules). Semantic equivalence looks only at the values of the expressions and not at their structure; Syntactic equivalence looks only at the structure of the expressions and not the values they produce.
Now, on to answer the question of how to do this in Isabelle.
Defining the relation
The standard way is to define a datatype for your formulas (I added some nicer infix syntax for it):
type_synonym vname = nat
datatype formula =
Atom vname
| FTrue
| Neg formula
| Conj formula formula (infixl "and" 60)
| Disj formula formula (infixl "or" 50)
definition "FFalse = Neg FTrue"
Then you can define the concept of ‘evaluating’ such a formula w.r.t. a given variable valuation:
primrec eval_formula :: "(vname ⇒ bool) ⇒ formula ⇒ bool" where
"eval_formula s (Atom x) ⟷ s x"
| "eval_formula _ FTrue ⟷ True"
| "eval_formula s (Neg a) ⟷ ¬eval_formula s a"
| "eval_formula s (a and b) ⟷ eval_formula s a ∧ eval_formula s b"
| "eval_formula s (a or b) ⟷ eval_formula s a ∨ eval_formula s b"
lemma eval_formula_False [simp]: "eval_formula s FFalse = False"
by (simp add: FFalse_def)
And, building on that, you can define the concept of semantic equivalence: two formulas are semantically equivalent if they evaluate to the same thing for all valuations:
definition formula_equiv_sem :: "formula ⇒ formula ⇒ bool" (infixl "≈" 40) where
"a ≈ b ⟷ (∀s. eval_formula s a = eval_formula s b)"
From your question, I gather that what you want to do is to define some kind of equivalence relation based on rewrite rules: two formulas are syntactically equivalent if you can transform one into another by applying some set of given rewrite rules.
That can be done e.g. with Isabelle's package for inductive predicates:
inductive formula_equiv :: "formula ⇒ formula ⇒ bool" (infixl "∼" 40) where
formula_refl [simp]: "a ∼ a"
| formula_sym: "a ∼ b ⟹ b ∼ a"
| formula_trans [trans]: "a ∼ b ⟹ b ∼ c ⟹ a ∼ c"
| neg_cong: "a ∼ b ⟹ Neg a ∼ Neg b"
| conj_cong: "a1 ∼ b1 ⟹ a2 ∼ b2 ⟹ a1 and a2 ∼ b1 and b2"
| disj_cong: "a1 ∼ b1 ⟹ a2 ∼ b2 ⟹ a1 or a2 ∼ b1 or b2"
| conj_commute: "a and b ∼ b and a"
| disj_commute: "a or b ∼ b or a"
| conj_assoc: "(a and b) and c ∼ a and (b and c)"
| disj_assoc: "(a or b) or c ∼ a or (b or c)"
| disj_conj: "a or (b and c) ∼ (a or b) and (a or c)"
| conj_disj: "a and (b or c) ∼ (a and b) or (a and c)"
| de_morgan1: "Neg (a and b) ∼ Neg a or Neg b"
| de_morgan2: "Neg (a or b) ∼ Neg a and Neg b"
| neg_neg: "Neg (Neg a) ∼ a"
| tnd: "a or Neg a ∼ FTrue"
| contr: "a and Neg a ∼ FFalse"
| disj_idem: "a or a ∼ a"
| conj_idem: "a and a ∼ a"
| conj_True: "a and FTrue ∼ a"
| disj_True: "a or FTrue ∼ FTrue"
The first six rules essentially set up rewriting (you can rewrite anything to itself, you can rewrite from left to right or from right to left, you can chain rewrite steps, if you can rewrite subterms, you can also rewrite the full term). The remaining rules are a few example rules that you may want to have in there (with no claim of completeness).
For more information on inductive predicates, refer to the manual of the predicate tool.
Proving stuff about it
So what can you do with this? Well, I expect that you will want to show that this is sound, i.e. that two formulas that are syntactically equivalent are also semantically equivalent:
lemma formula_equiv_syntactic_imp_semantic:
"a ∼ b ⟹ a ≈ b"
by (induction a b rule: formula_equiv.induct)
(auto simp: formula_equiv_sem_def)
You might also want to prove a few derived syntactical rules. For this, it is useful to have some convenient transitivity rules and setup the simplifier with the congruence rules:
lemmas formula_congs [simp] = neg_cong conj_cong disj_cong
lemma formula_trans_cong1 [trans]:
"a ∼ f b ⟹ b ∼ c ⟹ (⋀x y. x ∼ y ⟹ f x ∼ f y) ⟹ a ∼ f c"
by (rule formula_trans) simp_all
lemma formula_trans_cong2 [trans]:
"a ∼ b ⟹ f b ∼ f c ⟹ (⋀x y. x ∼ y ⟹ f x ∼ f y) ⟹ f a ∼ f c"
by (rule formula_trans) simp_all
Then we can do proofs like this:
lemma conj_False: "a and FFalse ∼ FFalse"
proof -
have "a and FFalse ∼ Neg (Neg (a and FFalse))"
by (rule formula_sym, rule neg_neg)
also have "Neg (a and FFalse) ∼ Neg a or Neg FFalse"
by (rule de_morgan1)
also have "Neg FFalse ∼ FTrue" unfolding FFalse_def by (rule neg_neg)
also have "Neg a or FTrue ∼ FTrue" by (rule disj_True)
also have "Neg FTrue = FFalse" unfolding FFalse_def ..
finally show ?thesis by - simp
qed
lemma disj_False: "a or FFalse ∼ a"
proof -
have "a or FFalse ∼ Neg (Neg (a or FFalse))" by (rule formula_sym, rule neg_neg)
also have "Neg (a or FFalse) ∼ Neg a and Neg FFalse" by (rule de_morgan2)
also have "Neg FFalse ∼ FTrue" unfolding FFalse_def by (rule neg_neg)
also have "Neg a and FTrue ∼ Neg a" by (rule conj_True)
also have "Neg (Neg a) ∼ a" by (rule neg_neg)
finally show ?thesis by - simp
qed
One would, of course, also like to prove completeness, i.e. that any two formulas that are semantically equivalent are also syntactically equivalent. For that, I think you will need a few more rules and then the proof is pretty complicated.
Why not axiomatization?
You mentioned axiomatization, and you might ask why you were advised not to use this for this purpose. Well, one reason is that axiomatization allows you to introduce inconsistencies into the system. You might ‘define’ two things to be equivalent and also define something else at another place that implies that they are not equivalent and then you have derive False and break everything. With the inductive predicate package, that cannot happen, because it proves automatically that your definitions are consistent. (by restricting them to be monotonic)
A more practical reason is that, as you can see above, you can do induction over an inductive predication, i.e. if you have two formulas that are synactically equivalent, you can do induction over the proof tree of their syntactic equivalence. In particular, you know that if two formulas are syntactically equivalent, that must be provable from the rules that you specified. If you just do axiomatization, you have no such guarantee – there may be many more syntactically equivalent formulas; with axiomatization, you could not even disprove something like Atom 0 ≈ Atom 1, unless you axiomatize something like that as well, which will be very ugly and very prone to accidental inconsistencies.
It is very rare for an Isabelle user to use axiomatization. I have been working with and on Isabelle for years and I have never used axiomatization. It is a very low-level feature designed to set up the basic underlying logic and much work has been invested in high-level definitional tools like typedef, datatype, fun, inductive, codatatype etc. to offer a definitional interface to the user that (hopefully) guarantee consistency to the user.
Appendix: Deciding semantic equivalence
If you're interested in using code generation to do interesting things on formulas: We can even decide semantic equivalence: we can simply test if the two expressions evaluate to the same results on the set of variables they containt. (syntactic equivalence is possible as well, but more difficult because you will have to get the inductive predicate package to compile usable code for it, and I did not manage to do that)
primrec vars :: "formula ⇒ vname list" where
"vars (Atom x) = [x]"
| "vars FTrue = []"
| "vars (Neg a) = vars a"
| "vars (Conj a b) = vars a # vars b"
| "vars (Disj a b) = vars a # vars b"
lemma eval_formula_cong:
"(⋀x. x ∈ set (vars a) ⟹ s x = s' x) ⟹ eval_formula s a = eval_formula s' a"
by (induction a) simp_all
primrec valuations :: "vname list ⇒ (vname ⇒ bool) list" where
"valuations [] = [λ_. False]"
| "valuations (x#xs) = [f' . f ← valuations xs, f' ← [f, fun_upd f x True]]"
lemma set_valuations: "set (valuations xs) = {f. ∀x. x∉set xs ⟶ f x = False}"
proof safe
case (goal2 f)
thus ?case
proof (induction xs arbitrary: f)
case (Cons x xs)
def f' ≡ "fun_upd f x False"
from Cons.prems have f': "f' ∈ set (valuations xs)"
by (intro Cons) (auto simp: f'_def)
show ?case
proof (cases "f x")
case False
hence "f' = f" by (intro ext) (simp add: f'_def)
with f' show ?thesis by auto
next
case True
hence "fun_upd f' x True = f" by (intro ext) (simp add: f'_def)
with f' show ?thesis by auto
qed
qed auto
qed (induction xs, auto)
lemma formula_equiv_sem_code [code]:
"a ≈ b ⟷ (∀s∈set (valuations (remdups (vars a # vars b))).
eval_formula s a = eval_formula s b)"
unfolding formula_equiv_sem_def
proof (rule iffI; rule ballI allI)
case (goal2 s)
def s' ≡ "λx. if x ∈ set (vars a # vars b) then s x else False"
have "s' ∈ set (valuations (remdups (vars a # vars b)))"
by (subst set_valuations) (auto simp: s'_def)
with goal2 have "eval_formula s' a = eval_formula s' b" by blast
also have "eval_formula s' a = eval_formula s a"
by (intro eval_formula_cong) (auto simp: s'_def)
also have "eval_formula s' b = eval_formula s b"
by (intro eval_formula_cong) (auto simp: s'_def)
finally show ?case .
qed auto
We can now simply ask Isabelle to compute whether two formulas are semantically equivalent:
value "Atom 0 and Atom 1 ≈ Atom 1 and Atom 0" (* True *)
value "Atom 0 and Atom 1 ≈ Atom 1 or Atom 0" (* False *)
You could even go further and write an automated proof method that decides a ≈ b for any formulas a and b by substituting all free variables with fresh atoms and then deciding the equivalence of those formulas (e.g. decide a and FFalse ≈ FFalse by deciding the equivalent Atom 0 and FFalse ≈ FFalse).
I have a partially-defined operator (disj_union below) on sets that I would like to lift to a quotient type (natq). Morally, I think this should be ok, because it is always possible to find in the equivalence class some representative for which the operator is defined [*]. However, I cannot complete the proof that the lifted definition preserves the equivalence, because disj_union is only partially defined. In my theory file below, I propose one way I have found to define my disj_union operator, but I don't like it because it features lots of abs and Rep functions, and I think it would be hard to work with (right?).
What is a good way to define this kind of thing using quotients in Isabelle?
theory My_Theory imports
"~~/src/HOL/Library/Quotient_Set"
begin
(* A ∪-operator that is defined only on disjoint operands. *)
definition "X ∩ Y = {} ⟹ disj_union X Y ≡ X ∪ Y"
(* Two sets are equivalent if they have the same cardinality. *)
definition "card_eq X Y ≡ finite X ∧ finite Y ∧ card X = card Y"
(* Quotient sets of naturals by this equivalence. *)
quotient_type natq = "nat set" / partial: card_eq
proof (intro part_equivpI)
show "∃x. card_eq x x" by (metis card_eq_def finite.emptyI)
show "symp card_eq" by (metis card_eq_def symp_def)
show "transp card_eq" by (metis card_eq_def transp_def)
qed
(* I want to lift my disj_union operator to the natq type.
But I cannot complete the proof, because disj_union is
only partially defined. *)
lift_definition natq_add :: "natq ⇒ natq ⇒ natq"
is "disj_union"
oops
(* Here is another attempt to define natq_add. I think it
is correct, but it looks hard to prove things about,
because it uses abstraction and representation functions
explicitly. *)
definition natq_add :: "natq ⇒ natq ⇒ natq"
where "natq_add X Y ≡
let (X',Y') = SOME (X',Y').
X' ∈ Rep_natq X ∧ Y' ∈ Rep_natq Y ∧ X' ∩ Y' = {}
in abs_natq (disj_union X' Y')"
end
[*] This is a little bit like how capture-avoiding substitution is only defined on the condition that bound variables do not clash; a condition that can always be satisfied by renaming to another representative in the alpha-equivalence class.
What about something like this (just an idea):
definition disj_union' :: "nat set ⇒ nat set ⇒ nat set"
where "disj_union' X Y ≡
let (X',Y') = SOME (X',Y').
card_eq X' X ∧ card_eq Y' Y ∧ X' ∩ Y' = {}
in disj_union X' Y'"
lift_definition natq_add :: "natq ⇒ natq ⇒ natq"
is "disj_union'" oops
For the record, here is Ondřej's suggestion (well, a slight amendment thereof, in which only one of the operands is renamed, not both) carried out to completion...
(* A version of disj_union that is always defined. *)
definition disj_union' :: "nat set ⇒ nat set ⇒ nat set"
where "disj_union' X Y ≡
let Y' = SOME Y'.
card_eq Y' Y ∧ X ∩ Y' = {}
in disj_union X Y'"
(* Can always choose a natural that is not in a given finite subset of ℕ. *)
lemma nats_infinite:
fixes A :: "nat set"
assumes "finite A"
shows "∃x. x ∉ A"
proof (rule ccontr, simp)
assume "∀x. x ∈ A"
hence "A = UNIV" by fast
hence "finite UNIV" using assms by fast
thus False by fast
qed
(* Can always choose n naturals that are not in a given finite subset of ℕ. *)
lemma nat_renaming:
fixes x :: "nat set" and n :: nat
assumes "finite x"
shows "∃z'. finite z' ∧ card z' = n ∧ x ∩ z' = {}"
using assms
apply (induct n)
apply (intro exI[of _ "{}"], simp)
apply (clarsimp)
apply (rule_tac x="insert (SOME y. y ∉ x ∪ z') z'" in exI)
apply (intro conjI, simp)
apply (rule someI2_ex, rule nats_infinite, simp, simp)+
done
lift_definition natq_add :: "natq ⇒ natq ⇒ natq"
is "disj_union'"
apply (unfold disj_union'_def card_eq_def)
apply (rule someI2_ex, simp add: nat_renaming)
apply (rule someI2_ex, simp add: nat_renaming)
apply (metis card.union_inter_neutral disj_union_def empty_iff finite_Un)
done