Mercurial
Mercurial
>
repos
>
isabelle
/ annotate
summary
|
shortlog
|
changelog
|
graph
|
tags
|
bookmarks
|
branches
|
files
|
changeset
|
file
|
latest
|
revisions
| annotate |
diff
|
comparison
|
raw
|
help
Find changesets by keywords (author, files, the commit message), revision number or hash, or
revset expression
.
src/FOLP/ex/if.thy
author
lcp
Wed, 03 May 1995 17:38:27 +0200
changeset 1105
136b05aa77ed
parent 0
a5a9c433f639
permissions
-rw-r--r--
Modified proofs for (q)split, fst, snd for new definitions. The rule f(q)splitE is now called (q)splitE and is weaker than before. The rule '(q)split' is now a meta-equality; this required modifying all proofs involving e.g. split RS trans.
Ignore whitespace changes -
Everywhere:
Within whitespace:
At end of lines:
0
a5a9c433f639
Initial revision
clasohm
parents:
diff
changeset
1
If = FOLP +
a5a9c433f639
Initial revision
clasohm
parents:
diff
changeset
2
consts if :: "[o,o,o]=>o"
a5a9c433f639
Initial revision
clasohm
parents:
diff
changeset
3
rules
a5a9c433f639
Initial revision
clasohm
parents:
diff
changeset
4
if_def "if(P,Q,R) == P&Q | ~P&R"
a5a9c433f639
Initial revision
clasohm
parents:
diff
changeset
5
end