author | blanchet |
Tue, 10 Jun 2014 11:38:53 +0200 | |
changeset 57199 | 472360558b22 |
parent 55945 | e96383acecf9 |
permissions | -rw-r--r-- |
55059 | 1 |
(* Title: HOL/BNF_Util.thy |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
2 |
Author: Dmitriy Traytel, TU Muenchen |
49282 | 3 |
Author: Jasmin Blanchette, TU Muenchen |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
4 |
Copyright 2012 |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
5 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
6 |
Library for bounded natural functors. |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
7 |
*) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
8 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
9 |
header {* Library for Bounded Natural Functors *} |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
10 |
|
49282 | 11 |
theory BNF_Util |
55056 | 12 |
imports BNF_Cardinal_Arithmetic |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
13 |
begin |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
14 |
|
55084 | 15 |
definition |
55945 | 16 |
rel_fun :: "('a \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'd \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('c \<Rightarrow> 'd) \<Rightarrow> bool" |
55084 | 17 |
where |
55945 | 18 |
"rel_fun A B = (\<lambda>f g. \<forall>x y. A x y \<longrightarrow> B (f x) (g y))" |
55084 | 19 |
|
55945 | 20 |
lemma rel_funI [intro]: |
55084 | 21 |
assumes "\<And>x y. A x y \<Longrightarrow> B (f x) (g y)" |
55945 | 22 |
shows "rel_fun A B f g" |
23 |
using assms by (simp add: rel_fun_def) |
|
55084 | 24 |
|
55945 | 25 |
lemma rel_funD: |
26 |
assumes "rel_fun A B f g" and "A x y" |
|
55084 | 27 |
shows "B (f x) (g y)" |
55945 | 28 |
using assms by (simp add: rel_fun_def) |
55084 | 29 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
30 |
definition collect where |
49463 | 31 |
"collect F x = (\<Union>f \<in> F. f x)" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
32 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
33 |
lemma fstI: "x = (y, z) \<Longrightarrow> fst x = y" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
34 |
by simp |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
35 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
36 |
lemma sndI: "x = (y, z) \<Longrightarrow> snd x = z" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
37 |
by simp |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
38 |
|
55066 | 39 |
lemma bijI': "\<lbrakk>\<And>x y. (f x = f y) = (x = y); \<And>y. \<exists>x. y = f x\<rbrakk> \<Longrightarrow> bij f" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
40 |
unfolding bij_def inj_on_def by auto blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
41 |
|
49312 | 42 |
(* Operator: *) |
43 |
definition "Gr A f = {(a, f a) | a. a \<in> A}" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
44 |
|
51893
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
49510
diff
changeset
|
45 |
definition "Grp A f = (\<lambda>a b. b = f a \<and> a \<in> A)" |
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
49510
diff
changeset
|
46 |
|
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55084
diff
changeset
|
47 |
definition vimage2p where |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55084
diff
changeset
|
48 |
"vimage2p f g R = (\<lambda>x y. R (f x) (g y))" |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55084
diff
changeset
|
49 |
|
55062 | 50 |
ML_file "Tools/BNF/bnf_util.ML" |
51 |
ML_file "Tools/BNF/bnf_tactics.ML" |
|
49283 | 52 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
53 |
end |